@@ -496,6 +496,10 @@ theorem eq_of_heq {α : Sort u} {a a' : α} (h : HEq a a') : Eq a a' :=
496496 h₁.rec (fun _ => rfl)
497497 this α α a a' h rfl
498498
499+ /-- Propositionally equal terms are also heterogeneously equal. -/
500+ theorem heq_of_eq (h : Eq a a') : HEq a a' :=
501+ Eq.subst h (HEq.refl a)
502+
499503/--
500504The product type, usually written `α × β`. Product types are also called pair or tuple types.
501505Elements of this type are pairs in which the first element is an `α` and the second element is a
@@ -2330,7 +2334,7 @@ def BitVec.decEq (x y : BitVec w) : Decidable (Eq x y) :=
23302334 | ⟨n⟩, ⟨m⟩ =>
23312335 dite (Eq n m)
23322336 (fun h => isTrue (h ▸ rfl))
2333- (fun h => isFalse (fun h' => BitVec.noConfusion h' (fun h' => absurd h' h)))
2337+ (fun h => isFalse (fun h' => BitVec.noConfusion rfl (heq_of_eq h') (fun h' => absurd (eq_of_heq h') h)))
23342338
23352339instance : DecidableEq (BitVec w) := BitVec.decEq
23362340
@@ -2921,15 +2925,15 @@ instance {α} : Inhabited (List α) where
29212925/-- Implements decidable equality for `List α`, assuming `α` has decidable equality. -/
29222926protected def List.hasDecEq {α : Type u} [DecidableEq α] : (a b : List α) → Decidable (Eq a b)
29232927 | nil, nil => isTrue rfl
2924- | cons _ _, nil => isFalse (fun h => List.noConfusion h )
2925- | nil, cons _ _ => isFalse (fun h => List.noConfusion h )
2928+ | cons _ _, nil => isFalse (fun h => List.noConfusion rfl (heq_of_eq h) )
2929+ | nil, cons _ _ => isFalse (fun h => List.noConfusion rfl (heq_of_eq h) )
29262930 | cons a as, cons b bs =>
29272931 match decEq a b with
29282932 | isTrue hab =>
29292933 match List.hasDecEq as bs with
29302934 | isTrue habs => isTrue (hab ▸ habs ▸ rfl)
2931- | isFalse nabs => isFalse (fun h => List.noConfusion h (fun _ habs => absurd habs nabs))
2932- | isFalse nab => isFalse (fun h => List.noConfusion h (fun hab _ => absurd hab nab))
2935+ | isFalse nabs => isFalse (fun h => List.noConfusion rfl (heq_of_eq h) (fun _ habs => absurd (eq_of_heq habs) nabs))
2936+ | isFalse nab => isFalse (fun h => List.noConfusion rfl (heq_of_eq h) (fun hab _ => absurd (eq_of_heq hab) nab))
29332937
29342938instance {α : Type u} [DecidableEq α] : DecidableEq (List α) := fun xs ys =>
29352939 /-
@@ -2939,32 +2943,32 @@ instance {α : Type u} [DecidableEq α] : DecidableEq (List α) := fun xs ys =>
29392943 match xs with
29402944 | .nil => match ys with
29412945 | .nil => isTrue rfl
2942- | .cons _ _ => isFalse List.noConfusion
2946+ | .cons _ _ => isFalse ( fun h => List.noConfusion rfl (heq_of_eq h))
29432947 | .cons a as => match ys with
2944- | .nil => isFalse List.noConfusion
2948+ | .nil => isFalse ( fun h => List.noConfusion rfl (heq_of_eq h))
29452949 | .cons b bs =>
29462950 match decEq a b with
29472951 | isTrue hab =>
29482952 match List.hasDecEq as bs with
29492953 | isTrue habs => isTrue (hab ▸ habs ▸ rfl)
2950- | isFalse nabs => isFalse (List.noConfusion · ( fun _ habs => absurd habs nabs))
2951- | isFalse nab => isFalse (List.noConfusion · ( fun hab _ => absurd hab nab))
2954+ | isFalse nabs => isFalse (fun h => List.noConfusion rfl (heq_of_eq h) ( fun _ habs => absurd (eq_of_heq habs) nabs))
2955+ | isFalse nab => isFalse (fun h => List.noConfusion rfl (heq_of_eq h) ( fun hab _ => absurd (eq_of_heq hab) nab))
29522956
29532957/--
29542958Equality with `List.nil` is decidable even if the underlying type does not have decidable equality.
29552959-/
29562960instance List.instDecidableNilEq (a : List α) : Decidable (Eq List.nil a) :=
29572961 match a with
29582962 | .nil => isTrue rfl
2959- | .cons _ _ => isFalse List.noConfusion
2963+ | .cons _ _ => isFalse ( fun h => List.noConfusion rfl (heq_of_eq h))
29602964
29612965/--
29622966Equality with `List.nil` is decidable even if the underlying type does not have decidable equality.
29632967-/
29642968instance List.instDecidableEqNil (a : List α) : Decidable (Eq a List.nil) :=
29652969 match a with
29662970 | .nil => isTrue rfl
2967- | .cons _ _ => isFalse List.noConfusion
2971+ | .cons _ _ => isFalse ( fun h => List.noConfusion rfl (heq_of_eq h))
29682972
29692973/--
29702974The length of a list.
0 commit comments