Skip to content

Commit 46ca51f

Browse files
committed
Prelude fixes
1 parent acf2aac commit 46ca51f

File tree

7 files changed

+36
-34
lines changed

7 files changed

+36
-34
lines changed

src/Init/Core.lean

Lines changed: 4 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -939,9 +939,7 @@ theorem HEq.subst {p : (T : Sort u) → T → Prop} (h₁ : a ≍ b) (h₂ : p
939939
@[symm] theorem HEq.symm (h : a ≍ b) : b ≍ a :=
940940
h.rec (HEq.refl a)
941941

942-
/-- Propositionally equal terms are also heterogeneously equal. -/
943-
theorem heq_of_eq (h : a = a') : a ≍ a' :=
944-
Eq.subst h (HEq.refl a)
942+
945943

946944
/-- Heterogeneous equality is transitive. -/
947945
theorem HEq.trans (h₁ : a ≍ b) (h₂ : b ≍ c) : a ≍ c :=
@@ -1370,7 +1368,7 @@ instance {α : Type u} {p : α → Prop} [BEq α] [LawfulBEq α] : LawfulBEq {x
13701368
instance {α : Sort u} {p : α → Prop} [DecidableEq α] : DecidableEq {x : α // p x} :=
13711369
fun ⟨a, h₁⟩ ⟨b, h₂⟩ =>
13721370
if h : a = b then isTrue (by subst h; exact rfl)
1373-
else isFalse (fun h' => Subtype.noConfusion h' (fun h' => absurd h' h))
1371+
else isFalse (fun h' => Subtype.noConfusion rfl .rfl (heq_of_eq h') (fun h' => absurd (eq_of_heq h') h))
13741372

13751373
end Subtype
13761374

@@ -1429,8 +1427,8 @@ instance [DecidableEq α] [DecidableEq β] : DecidableEq (α × β) :=
14291427
| isTrue e₁ =>
14301428
match decEq b b' with
14311429
| isTrue e₂ => isTrue (e₁ ▸ e₂ ▸ rfl)
1432-
| isFalse n₂ => isFalse fun h => Prod.noConfusion h fun _ e₂' => absurd e₂' n₂
1433-
| isFalse n₁ => isFalse fun h => Prod.noConfusion h fun e₁' _ => absurd e₁' n₁
1430+
| isFalse n₂ => isFalse fun h => Prod.noConfusion rfl rfl (heq_of_eq h) fun _ e₂' => absurd (eq_of_heq e₂') n₂
1431+
| isFalse n₁ => isFalse fun h => Prod.noConfusion rfl rfl (heq_of_eq h) fun e₁' _ => absurd (eq_of_heq e₁') n₁
14341432

14351433
instance [BEq α] [BEq β] : BEq (α × β) where
14361434
beq := fun (a₁, b₁) (a₂, b₂) => a₁ == a₂ && b₁ == b₂

src/Init/Data/Array/DecidableEq.lean

Lines changed: 5 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -99,31 +99,31 @@ instance instDecidableEq [DecidableEq α] : DecidableEq (Array α) := fun xs ys
9999
| ⟨[]⟩ =>
100100
match ys with
101101
| ⟨[]⟩ => isTrue rfl
102-
| ⟨_ :: _⟩ => isFalse (Array.noConfusion · (List.noConfusion ·))
102+
| ⟨_ :: _⟩ => isFalse (fun h => Array.noConfusion rfl (heq_of_eq h) (fun h => List.noConfusion rfl h))
103103
| ⟨a :: as⟩ =>
104104
match ys with
105-
| ⟨[]⟩ => isFalse (Array.noConfusion · (List.noConfusion ·))
105+
| ⟨[]⟩ => isFalse (fun h => Array.noConfusion rfl (heq_of_eq h) (fun h => List.noConfusion rfl h))
106106
| ⟨b :: bs⟩ => instDecidableEqImpl ⟨a :: as⟩ ⟨b :: bs⟩
107107

108108
@[csimp]
109109
theorem instDecidableEq_csimp : @instDecidableEq = @instDecidableEqImpl :=
110110
Subsingleton.allEq _ _
111-
111+
112112
/--
113113
Equality with `#[]` is decidable even if the underlying type does not have decidable equality.
114114
-/
115115
instance instDecidableEqEmp (xs : Array α) : Decidable (xs = #[]) :=
116116
match xs with
117117
| ⟨[]⟩ => isTrue rfl
118-
| ⟨_ :: _⟩ => isFalse (Array.noConfusion · (List.noConfusion ·))
118+
| ⟨_ :: _⟩ => isFalse (fun h => Array.noConfusion rfl (heq_of_eq h) (fun h => List.noConfusion rfl h))
119119

120120
/--
121121
Equality with `#[]` is decidable even if the underlying type does not have decidable equality.
122122
-/
123123
instance instDecidableEmpEq (ys : Array α) : Decidable (#[] = ys) :=
124124
match ys with
125125
| ⟨[]⟩ => isTrue rfl
126-
| ⟨_ :: _⟩ => isFalse (Array.noConfusion · (List.noConfusion ·))
126+
| ⟨_ :: _⟩ => isFalse (fun h => Array.noConfusion rfl (heq_of_eq h) (fun h => List.noConfusion rfl h))
127127

128128
theorem beq_eq_decide [BEq α] (xs ys : Array α) :
129129
(xs == ys) = if h : xs.size = ys.size then

src/Init/Data/List/Basic.lean

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -301,7 +301,7 @@ Examples:
301301
def getLast : ∀ (as : List α), as ≠ [] → α
302302
| [], h => absurd rfl h
303303
| [a], _ => a
304-
| _::b::as, _ => getLast (b::as) (fun h => List.noConfusion h)
304+
| _::b::as, _ => getLast (b::as) (fun h => List.noConfusion rfl (heq_of_eq h))
305305

306306
/-! ### getLast? -/
307307

@@ -318,7 +318,7 @@ Examples:
318318
-/
319319
def getLast? : List α → Option α
320320
| [] => none
321-
| a::as => some (getLast (a::as) (fun h => List.noConfusion h))
321+
| a::as => some (getLast (a::as) (fun h => List.noConfusion rfl (heq_of_eq h)))
322322

323323
@[simp, grind =] theorem getLast?_nil : @getLast? α [] = none := rfl
324324

@@ -337,7 +337,7 @@ Examples:
337337
-/
338338
def getLastD : (as : List α) → (fallback : α) → α
339339
| [], a₀ => a₀
340-
| a::as, _ => getLast (a::as) (fun h => List.noConfusion h)
340+
| a::as, _ => getLast (a::as) (fun h => List.noConfusion rfl (heq_of_eq h))
341341

342342
-- These aren't `simp` lemmas since we always simplify `getLastD` in terms of `getLast?`.
343343
theorem getLastD_nil {a : α} : getLastD [] a = a := rfl

src/Init/Data/List/BasicAux.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -57,7 +57,7 @@ Examples:
5757
@[expose]
5858
def getLast! [Inhabited α] : List α → α
5959
| [] => panic! "empty list"
60-
| a::as => getLast (a::as) (fun h => List.noConfusion h)
60+
| a::as => getLast (a::as) (fun h => List.noConfusion rfl (heq_of_eq h))
6161

6262
/-! ## Head and tail -/
6363

src/Init/Data/Option/Basic.lean

Lines changed: 5 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -22,12 +22,12 @@ instance instDecidableEq {α} [inst : DecidableEq α] : DecidableEq (Option α)
2222
match a with
2323
| none => match b with
2424
| none => .isTrue rfl
25-
| some _ => .isFalse Option.noConfusion
25+
| some _ => .isFalse (fun h => Option.noConfusion rfl (heq_of_eq h))
2626
| some a => match b with
27-
| none => .isFalse Option.noConfusion
27+
| none => .isFalse (fun h => Option.noConfusion rfl (heq_of_eq h))
2828
| some b => match inst a b with
2929
| .isTrue h => .isTrue (h ▸ rfl)
30-
| .isFalse n => .isFalse (Option.noConfusion · n)
30+
| .isFalse n => .isFalse (fun h => Option.noConfusion rfl (heq_of_eq h) (fun h' => absurd (eq_of_heq h') n))
3131

3232
/--
3333
Equality with `none` is decidable even if the wrapped type does not have decidable equality.
@@ -37,7 +37,7 @@ instance decidableEqNone (o : Option α) : Decidable (o = none) :=
3737
compatibility with the `DecidableEq` instance. -/
3838
match o with
3939
| none => .isTrue rfl
40-
| some _ => .isFalse Option.noConfusion
40+
| some _ => .isFalse (fun h => Option.noConfusion rfl (heq_of_eq h))
4141

4242
/--
4343
Equality with `none` is decidable even if the wrapped type does not have decidable equality.
@@ -47,7 +47,7 @@ instance decidableNoneEq (o : Option α) : Decidable (none = o) :=
4747
compatibility with the `DecidableEq` instance. -/
4848
match o with
4949
| none => .isTrue rfl
50-
| some _ => .isFalse Option.noConfusion
50+
| some _ => .isFalse (fun h => Option.noConfusion rfl (heq_of_eq h))
5151

5252
deriving instance BEq for Option
5353

src/Init/Data/Option/Instances.lean

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -16,9 +16,9 @@ namespace Option
1616

1717
theorem eq_of_eq_some {α : Type u} : ∀ {x y : Option α}, (∀ z, x = some z ↔ y = some z) → x = y
1818
| none, none, _ => rfl
19-
| none, some z, h => Option.noConfusion ((h z).2 rfl)
20-
| some z, none, h => Option.noConfusion ((h z).1 rfl)
21-
| some _, some w, h => Option.noConfusion ((h w).2 rfl) (congrArg some)
19+
| none, some z, h => Option.noConfusion rfl (heq_of_eq ((h z).2 rfl))
20+
| some z, none, h => Option.noConfusion rfl (heq_of_eq ((h z).1 rfl))
21+
| some _, some w, h => Option.noConfusion rfl (heq_of_eq ((h w).2 rfl)) (fun h => congrArg some (eq_of_heq h))
2222

2323
theorem eq_none_of_isNone {α : Type u} : ∀ {o : Option α}, o.isNone → o = none
2424
| none, _ => rfl

src/Init/Prelude.lean

Lines changed: 15 additions & 11 deletions
Original file line numberDiff line numberDiff line change
@@ -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
/--
500504
The product type, usually written `α × β`. Product types are also called pair or tuple types.
501505
Elements 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

23352339
instance : DecidableEq (BitVec w) := BitVec.decEq
23362340

@@ -2921,15 +2925,15 @@ instance {α} : Inhabited (List α) where
29212925
/-- Implements decidable equality for `List α`, assuming `α` has decidable equality. -/
29222926
protected 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

29342938
instance {α : 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
/--
29542958
Equality with `List.nil` is decidable even if the underlying type does not have decidable equality.
29552959
-/
29562960
instance 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
/--
29622966
Equality with `List.nil` is decidable even if the underlying type does not have decidable equality.
29632967
-/
29642968
instance 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
/--
29702974
The length of a list.

0 commit comments

Comments
 (0)