feat(library/standard): add basic heq theorems
Signed-off-by: Leonardo de Moura <leonardo@microsoft.com>
This commit is contained in:
parent
a38dc76b37
commit
24540056c5
2 changed files with 36 additions and 2 deletions
|
@ -80,7 +80,7 @@ theorem eq_id {A : Type} (a : A) : (a = a) = true
|
|||
:= eqt_intro (refl a)
|
||||
|
||||
theorem heq_id {A : Type} (a : A) : (a == a) = true
|
||||
:= eqt_intro (heq_refl a)
|
||||
:= eqt_intro (hrefl a)
|
||||
|
||||
theorem not_or (a b : Bool) : (¬ (a ∨ b)) = (¬ a ∧ ¬ b)
|
||||
:= boolext
|
||||
|
|
|
@ -247,8 +247,42 @@ theorem heq_to_eq {A : Type} {a b : A} (H : a == b) : a = b
|
|||
calc a = cast w a : symm (cast_eq w a)
|
||||
... = b : Hw
|
||||
|
||||
theorem heq_refl {A : Type} (a : A) : a == a
|
||||
theorem hrefl {A : Type} (a : A) : a == a
|
||||
:= eq_to_heq (refl a)
|
||||
|
||||
theorem heqt_elim {a : Bool} (H : a == true) : a
|
||||
:= eqt_elim (heq_to_eq H)
|
||||
|
||||
opaque_hint (hiding cast)
|
||||
|
||||
theorem hsubst {A B : Type} {a : A} {b : B} {P : ∀ (T : Type), T → Bool} (H1 : a == b) (H2 : P A a) : P B b
|
||||
:= have Haux1 : ∀ H : A = A, P A (cast H a), from
|
||||
assume H : A = A, subst (symm (cast_eq H a)) H2,
|
||||
obtains (Heq : A = B) (Hw : cast Heq a = b), from H1,
|
||||
have Haux2 : P B (cast Heq a), from subst Heq Haux1 Heq,
|
||||
subst Hw Haux2
|
||||
|
||||
theorem hsymm {A B : Type} {a : A} {b : B} (H : a == b) : b == a
|
||||
:= hsubst H (hrefl a)
|
||||
|
||||
theorem htrans {A B C : Type} {a : A} {b : B} {c : C} (H1 : a == b) (H2 : b == c) : a == c
|
||||
:= hsubst H2 H1
|
||||
|
||||
theorem htrans_left {A B : Type} {a : A} {b c : B} (H1 : a == b) (H2 : b = c) : a == c
|
||||
:= htrans H1 (eq_to_heq H2)
|
||||
|
||||
theorem htrans_right {A C : Type} {a b : A} {c : C} (H1 : a = b) (H2 : b == c) : a == c
|
||||
:= htrans (eq_to_heq H1) H2
|
||||
|
||||
calc_trans htrans
|
||||
calc_trans htrans_left
|
||||
calc_trans htrans_right
|
||||
|
||||
theorem cast_heq {A B : Type} (H : A = B) (a : A) : cast H a == a
|
||||
:= have H1 : ∀ (H : A = A) (a : A), cast H a == a, from
|
||||
λ H a, eq_to_heq (cast_eq H a),
|
||||
subst H H1 H a
|
||||
|
||||
theorem cast_eq_to_heq {A B : Type} {a : A} {b : B} {H : A = B} (H1 : cast H a = b) : a == b
|
||||
:= calc a == cast H a : hsymm (cast_heq H a)
|
||||
... = b : H1
|
||||
|
|
Loading…
Reference in a new issue