refactor(library/standard): cleanup notation
Signed-off-by: Leonardo de Moura <leonardo@microsoft.com>
This commit is contained in:
parent
a450ad5a95
commit
abe1dbd7e0
5 changed files with 32 additions and 35 deletions
|
@ -116,6 +116,6 @@ using decidable
|
||||||
theorem decidable_eq [instance] (a b : bool) : decidable (a = b)
|
theorem decidable_eq [instance] (a b : bool) : decidable (a = b)
|
||||||
:= bool_rec
|
:= bool_rec
|
||||||
(bool_rec (inl (refl '0)) (inr b0_ne_b1) b)
|
(bool_rec (inl (refl '0)) (inr b0_ne_b1) b)
|
||||||
(bool_rec (inr (not_eq_symm b0_ne_b1)) (inl (refl '1)) b)
|
(bool_rec (inr (ne_symm b0_ne_b1)) (inl (refl '1)) b)
|
||||||
a
|
a
|
||||||
end
|
end
|
||||||
|
|
|
@ -23,13 +23,13 @@ theorem prop_complete_swapped (a : Prop) : a = false ∨ a = true
|
||||||
a
|
a
|
||||||
|
|
||||||
theorem not_true : (¬true) = false
|
theorem not_true : (¬true) = false
|
||||||
:= have aux : ¬ (¬true) = true, from
|
:= have aux : (¬true) ≠ true, from
|
||||||
assume H : (¬true) = true,
|
assume H : (¬true) = true,
|
||||||
absurd_not_true (H⁻¹ ▸ trivial),
|
absurd_not_true (H⁻¹ ▸ trivial),
|
||||||
resolve_right (prop_complete (¬true)) aux
|
resolve_right (prop_complete (¬true)) aux
|
||||||
|
|
||||||
theorem not_false : (¬false) = true
|
theorem not_false : (¬false) = true
|
||||||
:= have aux : ¬ (¬false) = false, from
|
:= have aux : (¬false) ≠ false, from
|
||||||
assume H : (¬false) = false,
|
assume H : (¬false) = false,
|
||||||
H ▸ not_false_trivial,
|
H ▸ not_false_trivial,
|
||||||
resolve_right (prop_complete_swapped (¬false)) aux
|
resolve_right (prop_complete_swapped (¬false)) aux
|
||||||
|
|
|
@ -196,11 +196,8 @@ theorem a_neq_a_elim {A : Type} {a : A} (H : a ≠ a) : false
|
||||||
theorem ne_irrefl {A : Type} {a : A} (H : a ≠ a) : false
|
theorem ne_irrefl {A : Type} {a : A} (H : a ≠ a) : false
|
||||||
:= H (refl a)
|
:= H (refl a)
|
||||||
|
|
||||||
theorem not_eq_symm {A : Type} {a b : A} (H : ¬ a = b) : ¬ b = a
|
|
||||||
:= assume H1 : b = a, H (H1⁻¹)
|
|
||||||
|
|
||||||
theorem ne_symm {A : Type} {a b : A} (H : a ≠ b) : b ≠ a
|
theorem ne_symm {A : Type} {a b : A} (H : a ≠ b) : b ≠ a
|
||||||
:= not_eq_symm H
|
:= assume H1 : b = a, H (H1⁻¹)
|
||||||
|
|
||||||
theorem eq_ne_trans {A : Type} {a b c : A} (H1 : a = b) (H2 : b ≠ c) : a ≠ c
|
theorem eq_ne_trans {A : Type} {a b c : A} (H1 : a = b) (H2 : b ≠ c) : a ≠ c
|
||||||
:= H1⁻¹ ▸ H2
|
:= H1⁻¹ ▸ H2
|
||||||
|
|
Loading…
Reference in a new issue