refactor(builtin/kernel): use iff instead of = for Booleans
Signed-off-by: Leonardo de Moura <leonardo@microsoft.com>
This commit is contained in:
parent
5453a8f2f1
commit
4dc98bc73b
13 changed files with 154 additions and 118 deletions
|
@ -21,12 +21,17 @@ infixr 30 ∨ : or
|
||||||
|
|
||||||
definition and (a b : Bool) := ¬ (a → ¬ b)
|
definition and (a b : Bool) := ¬ (a → ¬ b)
|
||||||
|
|
||||||
definition implies (a b : Bool) := a → b
|
|
||||||
|
|
||||||
infixr 35 && : and
|
infixr 35 && : and
|
||||||
infixr 35 /\ : and
|
infixr 35 /\ : and
|
||||||
infixr 35 ∧ : and
|
infixr 35 ∧ : and
|
||||||
|
|
||||||
|
definition implies (a b : Bool) := a → b
|
||||||
|
|
||||||
|
definition iff (a b : Bool) := a == b
|
||||||
|
|
||||||
|
infixr 25 <-> : iff
|
||||||
|
infixr 25 ⟷ : iff
|
||||||
|
|
||||||
-- The Lean parser has special treatment for the constant exists.
|
-- The Lean parser has special treatment for the constant exists.
|
||||||
-- It allows us to write
|
-- It allows us to write
|
||||||
-- exists x y : A, P x y and ∃ x y : A, P x y
|
-- exists x y : A, P x y and ∃ x y : A, P x y
|
||||||
|
@ -98,8 +103,8 @@ theorem imp_eq_trans {a b c : Bool} (H1 : a → b) (H2 : b = c) : a → c
|
||||||
theorem eq_imp_trans {a b c : Bool} (H1 : a = b) (H2 : b → c) : a → c
|
theorem eq_imp_trans {a b c : Bool} (H1 : a = b) (H2 : b → c) : a → c
|
||||||
:= assume Ha, H2 (H1 ◂ Ha)
|
:= assume Ha, H2 (H1 ◂ Ha)
|
||||||
|
|
||||||
theorem not_not_eq (a : Bool) : (¬ ¬ a) = a
|
theorem not_not_eq (a : Bool) : ¬ ¬ a ⟷ a
|
||||||
:= case (λ x, (¬ ¬ x) = x) trivial trivial a
|
:= case (λ x, ¬ ¬ x ⟷ x) trivial trivial a
|
||||||
|
|
||||||
theorem not_not_elim {a : Bool} (H : ¬ ¬ a) : a
|
theorem not_not_elim {a : Bool} (H : ¬ ¬ a) : a
|
||||||
:= (not_not_eq a) ◂ H
|
:= (not_not_eq a) ◂ H
|
||||||
|
@ -223,14 +228,14 @@ theorem eqf_intro {a : Bool} (H : ¬ a) : a = false
|
||||||
:= boolext (assume H1 : a, absurd H1 H)
|
:= boolext (assume H1 : a, absurd H1 H)
|
||||||
(assume H2 : false, false_elim a H2)
|
(assume H2 : false, false_elim a H2)
|
||||||
|
|
||||||
theorem neq_elim {A : TypeU} {a b : A} (H : a ≠ b) : (a = b) = false
|
theorem neq_elim {A : TypeU} {a b : A} (H : a ≠ b) : a = b ⟷ false
|
||||||
:= eqf_intro H
|
:= eqf_intro H
|
||||||
|
|
||||||
theorem or_comm (a b : Bool) : (a ∨ b) = (b ∨ a)
|
theorem or_comm (a b : Bool) : a ∨ b ⟷ b ∨ a
|
||||||
:= boolext (assume H, or_elim H (λ H1, or_intror b H1) (λ H2, or_introl H2 a))
|
:= boolext (assume H, or_elim H (λ H1, or_intror b H1) (λ H2, or_introl H2 a))
|
||||||
(assume H, or_elim H (λ H1, or_intror a H1) (λ H2, or_introl H2 b))
|
(assume H, or_elim H (λ H1, or_intror a H1) (λ H2, or_introl H2 b))
|
||||||
|
|
||||||
theorem or_assoc (a b c : Bool) : ((a ∨ b) ∨ c) = (a ∨ (b ∨ c))
|
theorem or_assoc (a b c : Bool) : (a ∨ b) ∨ c ⟷ a ∨ (b ∨ c)
|
||||||
:= boolext (assume H : (a ∨ b) ∨ c,
|
:= boolext (assume H : (a ∨ b) ∨ c,
|
||||||
or_elim H (λ H1 : a ∨ b, or_elim H1 (λ Ha : a, or_introl Ha (b ∨ c))
|
or_elim H (λ H1 : a ∨ b, or_elim H1 (λ Ha : a, or_introl Ha (b ∨ c))
|
||||||
(λ Hb : b, or_intror a (or_introl Hb c)))
|
(λ Hb : b, or_intror a (or_introl Hb c)))
|
||||||
|
@ -240,120 +245,120 @@ theorem or_assoc (a b c : Bool) : ((a ∨ b) ∨ c) = (a ∨ (b ∨ c))
|
||||||
(λ H1 : b ∨ c, or_elim H1 (λ Hb : b, or_introl (or_intror a Hb) c)
|
(λ H1 : b ∨ c, or_elim H1 (λ Hb : b, or_introl (or_intror a Hb) c)
|
||||||
(λ Hc : c, or_intror (a ∨ b) Hc)))
|
(λ Hc : c, or_intror (a ∨ b) Hc)))
|
||||||
|
|
||||||
theorem or_id (a : Bool) : (a ∨ a) = a
|
theorem or_id (a : Bool) : a ∨ a ⟷ a
|
||||||
:= boolext (assume H, or_elim H (λ H1, H1) (λ H2, H2))
|
:= boolext (assume H, or_elim H (λ H1, H1) (λ H2, H2))
|
||||||
(assume H, or_introl H a)
|
(assume H, or_introl H a)
|
||||||
|
|
||||||
theorem or_falsel (a : Bool) : (a ∨ false) = a
|
theorem or_falsel (a : Bool) : a ∨ false ⟷ a
|
||||||
:= boolext (assume H, or_elim H (λ H1, H1) (λ H2, false_elim a H2))
|
:= boolext (assume H, or_elim H (λ H1, H1) (λ H2, false_elim a H2))
|
||||||
(assume H, or_introl H false)
|
(assume H, or_introl H false)
|
||||||
|
|
||||||
theorem or_falser (a : Bool) : (false ∨ a) = a
|
theorem or_falser (a : Bool) : false ∨ a ⟷ a
|
||||||
:= (or_comm false a) ⋈ (or_falsel a)
|
:= (or_comm false a) ⋈ (or_falsel a)
|
||||||
|
|
||||||
theorem or_truel (a : Bool) : (true ∨ a) = true
|
theorem or_truel (a : Bool) : true ∨ a ⟷ true
|
||||||
:= eqt_intro (case (λ x : Bool, true ∨ x) trivial trivial a)
|
:= eqt_intro (case (λ x : Bool, true ∨ x) trivial trivial a)
|
||||||
|
|
||||||
theorem or_truer (a : Bool) : (a ∨ true) = true
|
theorem or_truer (a : Bool) : a ∨ true ⟷ true
|
||||||
:= (or_comm a true) ⋈ (or_truel a)
|
:= (or_comm a true) ⋈ (or_truel a)
|
||||||
|
|
||||||
theorem or_tauto (a : Bool) : (a ∨ ¬ a) = true
|
theorem or_tauto (a : Bool) : a ∨ ¬ a ⟷ true
|
||||||
:= eqt_intro (em a)
|
:= eqt_intro (em a)
|
||||||
|
|
||||||
theorem and_comm (a b : Bool) : (a ∧ b) = (b ∧ a)
|
theorem and_comm (a b : Bool) : a ∧ b ⟷ b ∧ a
|
||||||
:= boolext (assume H, and_intro (and_elimr H) (and_eliml H))
|
:= boolext (assume H, and_intro (and_elimr H) (and_eliml H))
|
||||||
(assume H, and_intro (and_elimr H) (and_eliml H))
|
(assume H, and_intro (and_elimr H) (and_eliml H))
|
||||||
|
|
||||||
theorem and_id (a : Bool) : (a ∧ a) = a
|
theorem and_id (a : Bool) : a ∧ a ⟷ a
|
||||||
:= boolext (assume H, and_eliml H)
|
:= boolext (assume H, and_eliml H)
|
||||||
(assume H, and_intro H H)
|
(assume H, and_intro H H)
|
||||||
|
|
||||||
theorem and_assoc (a b c : Bool) : ((a ∧ b) ∧ c) = (a ∧ (b ∧ c))
|
theorem and_assoc (a b c : Bool) : (a ∧ b) ∧ c ⟷ a ∧ (b ∧ c)
|
||||||
:= boolext (assume H, and_intro (and_eliml (and_eliml H)) (and_intro (and_elimr (and_eliml H)) (and_elimr H)))
|
:= boolext (assume H, and_intro (and_eliml (and_eliml H)) (and_intro (and_elimr (and_eliml H)) (and_elimr H)))
|
||||||
(assume H, and_intro (and_intro (and_eliml H) (and_eliml (and_elimr H))) (and_elimr (and_elimr H)))
|
(assume H, and_intro (and_intro (and_eliml H) (and_eliml (and_elimr H))) (and_elimr (and_elimr H)))
|
||||||
|
|
||||||
theorem and_truer (a : Bool) : (a ∧ true) = a
|
theorem and_truer (a : Bool) : a ∧ true ⟷ a
|
||||||
:= boolext (assume H : a ∧ true, and_eliml H)
|
:= boolext (assume H : a ∧ true, and_eliml H)
|
||||||
(assume H : a, and_intro H trivial)
|
(assume H : a, and_intro H trivial)
|
||||||
|
|
||||||
theorem and_truel (a : Bool) : (true ∧ a) = a
|
theorem and_truel (a : Bool) : true ∧ a ⟷ a
|
||||||
:= trans (and_comm true a) (and_truer a)
|
:= trans (and_comm true a) (and_truer a)
|
||||||
|
|
||||||
theorem and_falsel (a : Bool) : (a ∧ false) = false
|
theorem and_falsel (a : Bool) : a ∧ false ⟷ false
|
||||||
:= boolext (assume H, and_elimr H)
|
:= boolext (assume H, and_elimr H)
|
||||||
(assume H, false_elim (a ∧ false) H)
|
(assume H, false_elim (a ∧ false) H)
|
||||||
|
|
||||||
theorem and_falser (a : Bool) : (false ∧ a) = false
|
theorem and_falser (a : Bool) : false ∧ a ⟷ false
|
||||||
:= (and_comm false a) ⋈ (and_falsel a)
|
:= (and_comm false a) ⋈ (and_falsel a)
|
||||||
|
|
||||||
theorem and_absurd (a : Bool) : (a ∧ ¬ a) = false
|
theorem and_absurd (a : Bool) : a ∧ ¬ a ⟷ false
|
||||||
:= boolext (assume H, absurd (and_eliml H) (and_elimr H))
|
:= boolext (assume H, absurd (and_eliml H) (and_elimr H))
|
||||||
(assume H, false_elim (a ∧ ¬ a) H)
|
(assume H, false_elim (a ∧ ¬ a) H)
|
||||||
|
|
||||||
theorem imp_truer (a : Bool) : (a → true) = true
|
theorem imp_truer (a : Bool) : (a → true) ⟷ true
|
||||||
:= case (λ x, (x → true) = true) trivial trivial a
|
:= case (λ x, (x → true) ⟷ true) trivial trivial a
|
||||||
|
|
||||||
theorem imp_truel (a : Bool) : (true → a) = a
|
theorem imp_truel (a : Bool) : (true → a) ⟷ a
|
||||||
:= case (λ x, (true → x) = x) trivial trivial a
|
:= case (λ x, (true → x) ⟷ x) trivial trivial a
|
||||||
|
|
||||||
theorem imp_falser (a : Bool) : (a → false) = ¬ a
|
theorem imp_falser (a : Bool) : (a → false) ⟷ ¬ a
|
||||||
:= refl _
|
:= refl _
|
||||||
|
|
||||||
theorem imp_falsel (a : Bool) : (false → a) = true
|
theorem imp_falsel (a : Bool) : (false → a) ⟷ true
|
||||||
:= case (λ x, (false → x) = true) trivial trivial a
|
:= case (λ x, (false → x) ⟷ true) trivial trivial a
|
||||||
|
|
||||||
theorem not_and (a b : Bool) : (¬ (a ∧ b)) = (¬ a ∨ ¬ b)
|
theorem not_and (a b : Bool) : ¬ (a ∧ b) ⟷ ¬ a ∨ ¬ b
|
||||||
:= case (λ x, (¬ (x ∧ b)) = (¬ x ∨ ¬ b))
|
:= case (λ x, ¬ (x ∧ b) ⟷ ¬ x ∨ ¬ b)
|
||||||
(case (λ y, (¬ (true ∧ y)) = (¬ true ∨ ¬ y)) trivial trivial b)
|
(case (λ y, ¬ (true ∧ y) ⟷ ¬ true ∨ ¬ y) trivial trivial b)
|
||||||
(case (λ y, (¬ (false ∧ y)) = (¬ false ∨ ¬ y)) trivial trivial b)
|
(case (λ y, ¬ (false ∧ y) ⟷ ¬ false ∨ ¬ y) trivial trivial b)
|
||||||
a
|
a
|
||||||
|
|
||||||
theorem not_and_elim {a b : Bool} (H : ¬ (a ∧ b)) : ¬ a ∨ ¬ b
|
theorem not_and_elim {a b : Bool} (H : ¬ (a ∧ b)) : ¬ a ∨ ¬ b
|
||||||
:= (not_and a b) ◂ H
|
:= (not_and a b) ◂ H
|
||||||
|
|
||||||
theorem not_or (a b : Bool) : (¬ (a ∨ b)) = (¬ a ∧ ¬ b)
|
theorem not_or (a b : Bool) : ¬ (a ∨ b) ⟷ ¬ a ∧ ¬ b
|
||||||
:= case (λ x, (¬ (x ∨ b)) = (¬ x ∧ ¬ b))
|
:= case (λ x, ¬ (x ∨ b) ⟷ ¬ x ∧ ¬ b)
|
||||||
(case (λ y, (¬ (true ∨ y)) = (¬ true ∧ ¬ y)) trivial trivial b)
|
(case (λ y, ¬ (true ∨ y) ⟷ ¬ true ∧ ¬ y) trivial trivial b)
|
||||||
(case (λ y, (¬ (false ∨ y)) = (¬ false ∧ ¬ y)) trivial trivial b)
|
(case (λ y, ¬ (false ∨ y) ⟷ ¬ false ∧ ¬ y) trivial trivial b)
|
||||||
a
|
a
|
||||||
|
|
||||||
theorem not_or_elim {a b : Bool} (H : ¬ (a ∨ b)) : ¬ a ∧ ¬ b
|
theorem not_or_elim {a b : Bool} (H : ¬ (a ∨ b)) : ¬ a ∧ ¬ b
|
||||||
:= (not_or a b) ◂ H
|
:= (not_or a b) ◂ H
|
||||||
|
|
||||||
theorem not_iff (a b : Bool) : (¬ (a = b)) = ((¬ a) = b)
|
theorem not_iff (a b : Bool) : ¬ (a ⟷ b) ⟷ (¬ a ⟷ b)
|
||||||
:= case (λ x, (¬ (x = b)) = ((¬ x) = b))
|
:= case (λ x, ¬ (x ⟷ b) ⟷ ((¬ x) ⟷ b))
|
||||||
(case (λ y, (¬ (true = y)) = ((¬ true) = y)) trivial trivial b)
|
(case (λ y, ¬ (true ⟷ y) ⟷ ((¬ true) ⟷ y)) trivial trivial b)
|
||||||
(case (λ y, (¬ (false = y)) = ((¬ false) = y)) trivial trivial b)
|
(case (λ y, ¬ (false ⟷ y) ⟷ ((¬ false) ⟷ y)) trivial trivial b)
|
||||||
a
|
a
|
||||||
|
|
||||||
theorem not_iff_elim {a b : Bool} (H : ¬ (a = b)) : (¬ a) = b
|
theorem not_iff_elim {a b : Bool} (H : ¬ (a ⟷ b)) : (¬ a) ⟷ b
|
||||||
:= (not_iff a b) ◂ H
|
:= (not_iff a b) ◂ H
|
||||||
|
|
||||||
theorem not_implies (a b : Bool) : (¬ (a → b)) = (a ∧ ¬ b)
|
theorem not_implies (a b : Bool) : ¬ (a → b) ⟷ a ∧ ¬ b
|
||||||
:= case (λ x, (¬ (x → b)) = (x ∧ ¬ b))
|
:= case (λ x, ¬ (x → b) ⟷ x ∧ ¬ b)
|
||||||
(case (λ y, (¬ (true → y)) = (true ∧ ¬ y)) trivial trivial b)
|
(case (λ y, ¬ (true → y) ⟷ true ∧ ¬ y) trivial trivial b)
|
||||||
(case (λ y, (¬ (false → y)) = (false ∧ ¬ y)) trivial trivial b)
|
(case (λ y, ¬ (false → y) ⟷ false ∧ ¬ y) trivial trivial b)
|
||||||
a
|
a
|
||||||
|
|
||||||
theorem not_implies_elim {a b : Bool} (H : ¬ (a → b)) : a ∧ ¬ b
|
theorem not_implies_elim {a b : Bool} (H : ¬ (a → b)) : a ∧ ¬ b
|
||||||
:= (not_implies a b) ◂ H
|
:= (not_implies a b) ◂ H
|
||||||
|
|
||||||
theorem not_congr {a b : Bool} (H : a = b) : (¬ a) = (¬ b)
|
theorem not_congr {a b : Bool} (H : a ⟷ b) : ¬ a ⟷ ¬ b
|
||||||
:= congr2 not H
|
:= congr2 not H
|
||||||
|
|
||||||
theorem eq_exists_intro {A : (Type U)} {P Q : A → Bool} (H : ∀ x : A, P x = Q x) : (∃ x : A, P x) = (∃ x : A, Q x)
|
theorem eq_exists_intro {A : (Type U)} {P Q : A → Bool} (H : ∀ x : A, P x ⟷ Q x) : (∃ x : A, P x) ⟷ (∃ x : A, Q x)
|
||||||
:= congr2 (Exists A) (funext H)
|
:= congr2 (Exists A) (funext H)
|
||||||
|
|
||||||
theorem not_forall (A : (Type U)) (P : A → Bool) : (¬ (∀ x : A, P x)) = (∃ x : A, ¬ P x)
|
theorem not_forall (A : (Type U)) (P : A → Bool) : ¬ (∀ x : A, P x) ⟷ (∃ x : A, ¬ P x)
|
||||||
:= calc (¬ ∀ x : A, P x) = (¬ ∀ x : A, ¬ ¬ P x) : not_congr (allext (λ x : A, symm (not_not_eq (P x))))
|
:= calc (¬ ∀ x : A, P x) = ¬ ∀ x : A, ¬ ¬ P x : not_congr (allext (λ x : A, symm (not_not_eq (P x))))
|
||||||
... = (∃ x : A, ¬ P x) : refl (∃ x : A, ¬ P x)
|
... = ∃ x : A, ¬ P x : refl (∃ x : A, ¬ P x)
|
||||||
|
|
||||||
theorem not_forall_elim {A : (Type U)} {P : A → Bool} (H : ¬ (∀ x : A, P x)) : ∃ x : A, ¬ P x
|
theorem not_forall_elim {A : (Type U)} {P : A → Bool} (H : ¬ (∀ x : A, P x)) : ∃ x : A, ¬ P x
|
||||||
:= (not_forall A P) ◂ H
|
:= (not_forall A P) ◂ H
|
||||||
|
|
||||||
theorem not_exists (A : (Type U)) (P : A → Bool) : (¬ ∃ x : A, P x) = (∀ x : A, ¬ P x)
|
theorem not_exists (A : (Type U)) (P : A → Bool) : ¬ (∃ x : A, P x) ⟷ (∀ x : A, ¬ P x)
|
||||||
:= calc (¬ ∃ x : A, P x) = (¬ ¬ ∀ x : A, ¬ P x) : refl (¬ ∃ x : A, P x)
|
:= calc (¬ ∃ x : A, P x) = ¬ ¬ ∀ x : A, ¬ P x : refl (¬ ∃ x : A, P x)
|
||||||
... = (∀ x : A, ¬ P x) : not_not_eq (∀ x : A, ¬ P x)
|
... = ∀ x : A, ¬ P x : not_not_eq (∀ x : A, ¬ P x)
|
||||||
|
|
||||||
theorem not_exists_elim {A : (Type U)} {P : A → Bool} (H : ¬ ∃ x : A, P x) : ∀ x : A, ¬ P x
|
theorem not_exists_elim {A : (Type U)} {P : A → Bool} (H : ¬ ∃ x : A, P x) : ∀ x : A, ¬ P x
|
||||||
:= (not_exists A P) ◂ H
|
:= (not_exists A P) ◂ H
|
||||||
|
@ -373,7 +378,7 @@ theorem exists_unfold2 {A : TypeU} {P : A → Bool} (a : A) (H : P a ∨ (∃ x
|
||||||
(λ (w : A) (Hw : w ≠ a ∧ P w),
|
(λ (w : A) (Hw : w ≠ a ∧ P w),
|
||||||
exists_intro w (and_elimr Hw)))
|
exists_intro w (and_elimr Hw)))
|
||||||
|
|
||||||
theorem exists_unfold {A : TypeU} (P : A → Bool) (a : A) : (∃ x : A, P x) = (P a ∨ (∃ x : A, x ≠ a ∧ P x))
|
theorem exists_unfold {A : TypeU} (P : A → Bool) (a : A) : (∃ x : A, P x) ⟷ (P a ∨ (∃ x : A, x ≠ a ∧ P x))
|
||||||
:= boolext (assume H : (∃ x : A, P x), exists_unfold1 a H)
|
:= boolext (assume H : (∃ x : A, P x), exists_unfold1 a H)
|
||||||
(assume H : (P a ∨ (∃ x : A, x ≠ a ∧ P x)), exists_unfold2 a H)
|
(assume H : (P a ∨ (∃ x : A, x ≠ a ∧ P x)), exists_unfold2 a H)
|
||||||
|
|
||||||
|
@ -385,10 +390,10 @@ theorem left_comm {A : TypeU} {R : A -> A -> A} (comm : ∀ x y, R x y = R y x)
|
||||||
... = R (R y x) z : { comm x y }
|
... = R (R y x) z : { comm x y }
|
||||||
... = R y (R x z) : assoc y x z
|
... = R y (R x z) : assoc y x z
|
||||||
|
|
||||||
theorem and_left_comm (a b c : Bool) : (a ∧ (b ∧ c)) = (b ∧ (a ∧ c))
|
theorem and_left_comm (a b c : Bool) : a ∧ (b ∧ c) ⟷ b ∧ (a ∧ c)
|
||||||
:= left_comm and_comm and_assoc a b c
|
:= left_comm and_comm and_assoc a b c
|
||||||
|
|
||||||
theorem or_left_comm (a b c : Bool) : (a ∨ (b ∨ c)) = (b ∨ (a ∨ c))
|
theorem or_left_comm (a b c : Bool) : a ∨ (b ∨ c) ⟷ b ∨ (a ∨ c)
|
||||||
:= left_comm or_comm or_assoc a b c
|
:= left_comm or_comm or_assoc a b c
|
||||||
|
|
||||||
-- Congruence theorems for contextual simplification
|
-- Congruence theorems for contextual simplification
|
||||||
|
@ -458,24 +463,24 @@ theorem imp_congr {a b c d : Bool} (H_ac : a = c) (H_bd : ∀ (H_c : c), b = d)
|
||||||
:= imp_congrr (λ H, H_ac) H_bd
|
:= imp_congrr (λ H, H_ac) H_bd
|
||||||
|
|
||||||
-- In the following theorems we are using the fact that a ∨ b is defined as ¬ a → b
|
-- In the following theorems we are using the fact that a ∨ b is defined as ¬ a → b
|
||||||
theorem or_congrr {a b c d : Bool} (H_ac : ∀ (H_nb : ¬ b), a = c) (H_bd : ∀ (H_nc : ¬ c), b = d) : (a ∨ b) = (c ∨ d)
|
theorem or_congrr {a b c d : Bool} (H_ac : ∀ (H_nb : ¬ b), a = c) (H_bd : ∀ (H_nc : ¬ c), b = d) : a ∨ b ⟷ c ∨ d
|
||||||
:= imp_congrr (λ H_nb : ¬ b, congr2 not (H_ac H_nb)) H_bd
|
:= imp_congrr (λ H_nb : ¬ b, congr2 not (H_ac H_nb)) H_bd
|
||||||
theorem or_congrl {a b c d : Bool} (H_ac : ∀ (H_nd : ¬ d), a = c) (H_bd : ∀ (H_na : ¬ a), b = d) : (a ∨ b) = (c ∨ d)
|
theorem or_congrl {a b c d : Bool} (H_ac : ∀ (H_nd : ¬ d), a = c) (H_bd : ∀ (H_na : ¬ a), b = d) : a ∨ b ⟷ c ∨ d
|
||||||
:= imp_congrl (λ H_nd : ¬ d, congr2 not (H_ac H_nd)) H_bd
|
:= imp_congrl (λ H_nd : ¬ d, congr2 not (H_ac H_nd)) H_bd
|
||||||
-- (Common case) simplify a to c, and then b to d using the fact that ¬ c is true
|
-- (Common case) simplify a to c, and then b to d using the fact that ¬ c is true
|
||||||
theorem or_congr {a b c d : Bool} (H_ac : a = c) (H_bd : ∀ (H_nc : ¬ c), b = d) : (a ∨ b) = (c ∨ d)
|
theorem or_congr {a b c d : Bool} (H_ac : a = c) (H_bd : ∀ (H_nc : ¬ c), b = d) : a ∨ b ⟷ c ∨ d
|
||||||
:= or_congrr (λ H, H_ac) H_bd
|
:= or_congrr (λ H, H_ac) H_bd
|
||||||
|
|
||||||
-- In the following theorems we are using the fact hat a ∧ b is defined as ¬ (a → ¬ b)
|
-- In the following theorems we are using the fact hat a ∧ b is defined as ¬ (a → ¬ b)
|
||||||
theorem and_congrr {a b c d : Bool} (H_ac : ∀ (H_b : b), a = c) (H_bd : ∀ (H_c : c), b = d) : (a ∧ b) = (c ∧ d)
|
theorem and_congrr {a b c d : Bool} (H_ac : ∀ (H_b : b), a = c) (H_bd : ∀ (H_c : c), b = d) : a ∧ b ⟷ c ∧ d
|
||||||
:= congr2 not (imp_congrr (λ (H_nnb : ¬ ¬ b), H_ac (not_not_elim H_nnb)) (λ H_c : c, congr2 not (H_bd H_c)))
|
:= congr2 not (imp_congrr (λ (H_nnb : ¬ ¬ b), H_ac (not_not_elim H_nnb)) (λ H_c : c, congr2 not (H_bd H_c)))
|
||||||
theorem and_congrl {a b c d : Bool} (H_ac : ∀ (H_d : d), a = c) (H_bd : ∀ (H_a : a), b = d) : (a ∧ b) = (c ∧ d)
|
theorem and_congrl {a b c d : Bool} (H_ac : ∀ (H_d : d), a = c) (H_bd : ∀ (H_a : a), b = d) : a ∧ b ⟷ c ∧ d
|
||||||
:= congr2 not (imp_congrl (λ (H_nnd : ¬ ¬ d), H_ac (not_not_elim H_nnd)) (λ H_a : a, congr2 not (H_bd H_a)))
|
:= congr2 not (imp_congrl (λ (H_nnd : ¬ ¬ d), H_ac (not_not_elim H_nnd)) (λ H_a : a, congr2 not (H_bd H_a)))
|
||||||
-- (Common case) simplify a to c, and then b to d using the fact that c is true
|
-- (Common case) simplify a to c, and then b to d using the fact that c is true
|
||||||
theorem and_congr {a b c d : Bool} (H_ac : a = c) (H_bd : ∀ (H_c : c), b = d) : (a ∧ b) = (c ∧ d)
|
theorem and_congr {a b c d : Bool} (H_ac : a = c) (H_bd : ∀ (H_c : c), b = d) : a ∧ b ⟷ c ∧ d
|
||||||
:= and_congrr (λ H, H_ac) H_bd
|
:= and_congrr (λ H, H_ac) H_bd
|
||||||
|
|
||||||
theorem forall_or_distributer {A : TypeU} (p : Bool) (φ : A → Bool) : (∀ x, p ∨ φ x) = (p ∨ ∀ x, φ x)
|
theorem forall_or_distributer {A : TypeU} (p : Bool) (φ : A → Bool) : (∀ x, p ∨ φ x) ⟷ p ∨ ∀ x, φ x
|
||||||
:= boolext
|
:= boolext
|
||||||
(assume H : (∀ x, p ∨ φ x),
|
(assume H : (∀ x, p ∨ φ x),
|
||||||
or_elim (em p)
|
or_elim (em p)
|
||||||
|
@ -488,19 +493,19 @@ theorem forall_or_distributer {A : TypeU} (p : Bool) (φ : A → Bool) : (∀ x,
|
||||||
(λ H1 : p, or_introl H1 (φ x))
|
(λ H1 : p, or_introl H1 (φ x))
|
||||||
(λ H2 : (∀ x, φ x), or_intror p (H2 x)))
|
(λ H2 : (∀ x, φ x), or_intror p (H2 x)))
|
||||||
|
|
||||||
theorem forall_or_distributel {A : TypeU} (p : Bool) (φ : A → Bool) : (∀ x, φ x ∨ p) = ((∀ x, φ x) ∨ p)
|
theorem forall_or_distributel {A : TypeU} (p : Bool) (φ : A → Bool) : (∀ x, φ x ∨ p) ⟷ (∀ x, φ x) ∨ p
|
||||||
:= calc (∀ x, φ x ∨ p) = (∀ x, p ∨ φ x) : allext (λ x, or_comm (φ x) p)
|
:= calc (∀ x, φ x ∨ p) = (∀ x, p ∨ φ x) : allext (λ x, or_comm (φ x) p)
|
||||||
... = (p ∨ ∀ x, φ x) : forall_or_distributer p φ
|
... = (p ∨ ∀ x, φ x) : forall_or_distributer p φ
|
||||||
... = ((∀ x, φ x) ∨ p) : or_comm p (∀ x, φ x)
|
... = ((∀ x, φ x) ∨ p) : or_comm p (∀ x, φ x)
|
||||||
|
|
||||||
theorem forall_and_distribute {A : TypeU} (φ ψ : A → Bool) : (∀ x, φ x ∧ ψ x) = ((∀ x, φ x) ∧ (∀ x, ψ x))
|
theorem forall_and_distribute {A : TypeU} (φ ψ : A → Bool) : (∀ x, φ x ∧ ψ x) ⟷ (∀ x, φ x) ∧ (∀ x, ψ x)
|
||||||
:= boolext
|
:= boolext
|
||||||
(assume H : (∀ x, φ x ∧ ψ x),
|
(assume H : (∀ x, φ x ∧ ψ x),
|
||||||
and_intro (take x, and_eliml (H x)) (take x, and_elimr (H x)))
|
and_intro (take x, and_eliml (H x)) (take x, and_elimr (H x)))
|
||||||
(assume H : (∀ x, φ x) ∧ (∀ x, ψ x),
|
(assume H : (∀ x, φ x) ∧ (∀ x, ψ x),
|
||||||
take x, and_intro (and_eliml H x) (and_elimr H x))
|
take x, and_intro (and_eliml H x) (and_elimr H x))
|
||||||
|
|
||||||
theorem exists_and_distributer {A : TypeU} (p : Bool) (φ : A → Bool) : (∃ x, p ∧ φ x) = (p ∧ ∃ x, φ x)
|
theorem exists_and_distributer {A : TypeU} (p : Bool) (φ : A → Bool) : (∃ x, p ∧ φ x) ⟷ p ∧ ∃ x, φ x
|
||||||
:= boolext
|
:= boolext
|
||||||
(assume H : (∃ x, p ∧ φ x),
|
(assume H : (∃ x, p ∧ φ x),
|
||||||
obtain (w : A) (Hw : p ∧ φ w), from H,
|
obtain (w : A) (Hw : p ∧ φ w), from H,
|
||||||
|
@ -509,12 +514,12 @@ theorem exists_and_distributer {A : TypeU} (p : Bool) (φ : A → Bool) : (∃ x
|
||||||
obtain (w : A) (Hw : φ w), from (and_elimr H),
|
obtain (w : A) (Hw : φ w), from (and_elimr H),
|
||||||
exists_intro w (and_intro (and_eliml H) Hw))
|
exists_intro w (and_intro (and_eliml H) Hw))
|
||||||
|
|
||||||
theorem exists_and_distributel {A : TypeU} (p : Bool) (φ : A → Bool) : (∃ x, φ x ∧ p) = ((∃ x, φ x) ∧ p)
|
theorem exists_and_distributel {A : TypeU} (p : Bool) (φ : A → Bool) : (∃ x, φ x ∧ p) ⟷ (∃ x, φ x) ∧ p
|
||||||
:= calc (∃ x, φ x ∧ p) = (∃ x, p ∧ φ x) : eq_exists_intro (λ x, and_comm (φ x) p)
|
:= calc (∃ x, φ x ∧ p) = (∃ x, p ∧ φ x) : eq_exists_intro (λ x, and_comm (φ x) p)
|
||||||
... = (p ∧ (∃ x, φ x)) : exists_and_distributer p φ
|
... = (p ∧ (∃ x, φ x)) : exists_and_distributer p φ
|
||||||
... = ((∃ x, φ x) ∧ p) : and_comm p (∃ x, φ x)
|
... = ((∃ x, φ x) ∧ p) : and_comm p (∃ x, φ x)
|
||||||
|
|
||||||
theorem exists_or_distribute {A : TypeU} (φ ψ : A → Bool) : (∃ x, φ x ∨ ψ x) = ((∃ x, φ x) ∨ (∃ x, ψ x))
|
theorem exists_or_distribute {A : TypeU} (φ ψ : A → Bool) : (∃ x, φ x ∨ ψ x) ⟷ (∃ x, φ x) ∨ (∃ x, ψ x)
|
||||||
:= boolext
|
:= boolext
|
||||||
(assume H : (∃ x, φ x ∨ ψ x),
|
(assume H : (∃ x, φ x ∨ ψ x),
|
||||||
obtain (w : A) (Hw : φ w ∨ ψ w), from H,
|
obtain (w : A) (Hw : φ w ∨ ψ w), from H,
|
||||||
|
|
Binary file not shown.
|
@ -28,7 +28,7 @@
|
||||||
'(("\\_<\\(Bool\\|Int\\|Nat\\|Real\\|Type\\|TypeU\\|ℕ\\|ℤ\\)\\_>" . 'font-lock-type-face)
|
'(("\\_<\\(Bool\\|Int\\|Nat\\|Real\\|Type\\|TypeU\\|ℕ\\|ℤ\\)\\_>" . 'font-lock-type-face)
|
||||||
("\\_<\\(calc\\|have\\|in\\|let\\|forall\\|exists\\|if\\|then\\|else\\|assume\\|take\\|obtain\\|from\\)\\_>" . font-lock-keyword-face)
|
("\\_<\\(calc\\|have\\|in\\|let\\|forall\\|exists\\|if\\|then\\|else\\|assume\\|take\\|obtain\\|from\\)\\_>" . font-lock-keyword-face)
|
||||||
("\"[^\"]*\"" . 'font-lock-string-face)
|
("\"[^\"]*\"" . 'font-lock-string-face)
|
||||||
("\\(->\\|/\\\\\\|==\\|\\\\/\\|[*+/<=>¬∧∨≠≤≥-]\\)" . 'font-lock-constant-face)
|
("\\(->\\|/\\\\\\|==\\|\\\\/\\|[*+/<=>¬∧∨≠≤≥-⟷]\\)" . 'font-lock-constant-face)
|
||||||
("\\(λ\\|→\\|∃\\|∀\\|:\\|:=\\)" . font-lock-constant-face)
|
("\\(λ\\|→\\|∃\\|∀\\|:\\|:=\\)" . font-lock-constant-face)
|
||||||
("\\_<\\(\\b.*_tac\\|Cond\\|OrElse\\|T\\(?:hen\\|ry\\)\\|When\\|apply\\|b\\(?:ack\\|eta\\)\\|done\\|exact\\)\\_>" . 'font-lock-constant-face)
|
("\\_<\\(\\b.*_tac\\|Cond\\|OrElse\\|T\\(?:hen\\|ry\\)\\|When\\|apply\\|b\\(?:ack\\|eta\\)\\|done\\|exact\\)\\_>" . 'font-lock-constant-face)
|
||||||
("\\(universe\\|theorem\\|axiom\\|definition\\|variable\\|builtin\\)[ \t]*\\([^ \t\n]*\\)" (2 'font-lock-function-name-face))
|
("\\(universe\\|theorem\\|axiom\\|definition\\|variable\\|builtin\\)[ \t]*\\([^ \t\n]*\\)" (2 'font-lock-function-name-face))
|
||||||
|
|
|
@ -12,6 +12,7 @@ MK_CONSTANT(not_fn, name("not"));
|
||||||
MK_CONSTANT(or_fn, name("or"));
|
MK_CONSTANT(or_fn, name("or"));
|
||||||
MK_CONSTANT(and_fn, name("and"));
|
MK_CONSTANT(and_fn, name("and"));
|
||||||
MK_CONSTANT(implies_fn, name("implies"));
|
MK_CONSTANT(implies_fn, name("implies"));
|
||||||
|
MK_CONSTANT(iff_fn, name("iff"));
|
||||||
MK_CONSTANT(exists_fn, name("exists"));
|
MK_CONSTANT(exists_fn, name("exists"));
|
||||||
MK_CONSTANT(eq_fn, name("eq"));
|
MK_CONSTANT(eq_fn, name("eq"));
|
||||||
MK_CONSTANT(neq_fn, name("neq"));
|
MK_CONSTANT(neq_fn, name("neq"));
|
||||||
|
|
|
@ -25,6 +25,10 @@ expr mk_implies_fn();
|
||||||
bool is_implies_fn(expr const & e);
|
bool is_implies_fn(expr const & e);
|
||||||
inline bool is_implies(expr const & e) { return is_app(e) && is_implies_fn(arg(e, 0)); }
|
inline bool is_implies(expr const & e) { return is_app(e) && is_implies_fn(arg(e, 0)); }
|
||||||
inline expr mk_implies(expr const & e1, expr const & e2) { return mk_app({mk_implies_fn(), e1, e2}); }
|
inline expr mk_implies(expr const & e1, expr const & e2) { return mk_app({mk_implies_fn(), e1, e2}); }
|
||||||
|
expr mk_iff_fn();
|
||||||
|
bool is_iff_fn(expr const & e);
|
||||||
|
inline bool is_iff(expr const & e) { return is_app(e) && is_iff_fn(arg(e, 0)); }
|
||||||
|
inline expr mk_iff(expr const & e1, expr const & e2) { return mk_app({mk_iff_fn(), e1, e2}); }
|
||||||
expr mk_exists_fn();
|
expr mk_exists_fn();
|
||||||
bool is_exists_fn(expr const & e);
|
bool is_exists_fn(expr const & e);
|
||||||
inline bool is_exists(expr const & e) { return is_app(e) && is_exists_fn(arg(e, 0)); }
|
inline bool is_exists(expr const & e) { return is_app(e) && is_exists_fn(arg(e, 0)); }
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
add_library(library kernel_bindings.cpp deep_copy.cpp
|
add_library(library kernel_bindings.cpp deep_copy.cpp
|
||||||
context_to_lambda.cpp placeholder.cpp expr_lt.cpp substitution.cpp
|
context_to_lambda.cpp placeholder.cpp expr_lt.cpp substitution.cpp
|
||||||
fo_unify.cpp bin_op.cpp eq_heq.cpp io_state_stream.cpp printer.cpp
|
fo_unify.cpp bin_op.cpp equality.cpp io_state_stream.cpp printer.cpp
|
||||||
hop_match.cpp ite.cpp)
|
hop_match.cpp ite.cpp)
|
||||||
|
|
||||||
target_link_libraries(library ${LEAN_LIBS})
|
target_link_libraries(library ${LEAN_LIBS})
|
||||||
|
|
|
@ -21,7 +21,7 @@ Author: Leonardo de Moura
|
||||||
#include "kernel/type_checker.h"
|
#include "kernel/type_checker.h"
|
||||||
#include "kernel/update_expr.h"
|
#include "kernel/update_expr.h"
|
||||||
#include "library/printer.h"
|
#include "library/printer.h"
|
||||||
#include "library/eq_heq.h"
|
#include "library/equality.h"
|
||||||
#include "library/elaborator/elaborator.h"
|
#include "library/elaborator/elaborator.h"
|
||||||
#include "library/elaborator/elaborator_justification.h"
|
#include "library/elaborator/elaborator_justification.h"
|
||||||
|
|
||||||
|
@ -1372,9 +1372,9 @@ class elaborator::imp {
|
||||||
r = process_metavar(c, b, a, false);
|
r = process_metavar(c, b, a, false);
|
||||||
if (r != Continue) { return r == Processed; }
|
if (r != Continue) { return r == Processed; }
|
||||||
|
|
||||||
if (is_eq_heq(a) && is_eq_heq(b)) {
|
if (is_equality(a) && is_equality(b)) {
|
||||||
expr_pair p1 = eq_heq_args(a);
|
expr_pair p1 = get_equality_args(a);
|
||||||
expr_pair p2 = eq_heq_args(b);
|
expr_pair p2 = get_equality_args(b);
|
||||||
justification new_jst(new destruct_justification(c));
|
justification new_jst(new destruct_justification(c));
|
||||||
push_new_eq_constraint(ctx, p1.first, p2.first, new_jst);
|
push_new_eq_constraint(ctx, p1.first, p2.first, new_jst);
|
||||||
push_new_eq_constraint(ctx, p1.second, p2.second, new_jst);
|
push_new_eq_constraint(ctx, p1.second, p2.second, new_jst);
|
||||||
|
|
|
@ -1,22 +0,0 @@
|
||||||
/*
|
|
||||||
Copyright (c) 2013 Microsoft Corporation. All rights reserved.
|
|
||||||
Released under Apache 2.0 license as described in the file LICENSE.
|
|
||||||
|
|
||||||
Author: Leonardo de Moura
|
|
||||||
*/
|
|
||||||
#include "library/expr_pair.h"
|
|
||||||
#include "kernel/kernel.h"
|
|
||||||
|
|
||||||
namespace lean {
|
|
||||||
bool is_eq_heq(expr const & e) {
|
|
||||||
return is_heq(e) || is_eq(e);
|
|
||||||
}
|
|
||||||
|
|
||||||
expr_pair eq_heq_args(expr const & e) {
|
|
||||||
lean_assert(is_heq(e) || is_eq(e));
|
|
||||||
if (is_heq(e))
|
|
||||||
return expr_pair(heq_lhs(e), heq_rhs(e));
|
|
||||||
else
|
|
||||||
return expr_pair(arg(e, 2), arg(e, 3));
|
|
||||||
}
|
|
||||||
}
|
|
44
src/library/equality.cpp
Normal file
44
src/library/equality.cpp
Normal file
|
@ -0,0 +1,44 @@
|
||||||
|
/*
|
||||||
|
Copyright (c) 2013 Microsoft Corporation. All rights reserved.
|
||||||
|
Released under Apache 2.0 license as described in the file LICENSE.
|
||||||
|
|
||||||
|
Author: Leonardo de Moura
|
||||||
|
*/
|
||||||
|
#include "library/expr_pair.h"
|
||||||
|
#include "kernel/kernel.h"
|
||||||
|
|
||||||
|
namespace lean {
|
||||||
|
bool is_equality(expr const & e) {
|
||||||
|
return is_eq(e) || is_iff(e) || is_heq(e);
|
||||||
|
}
|
||||||
|
|
||||||
|
bool is_equality(expr const & e, expr & lhs, expr & rhs) {
|
||||||
|
if (is_eq(e)) {
|
||||||
|
lhs = arg(e, 2);
|
||||||
|
rhs = arg(e, 3);
|
||||||
|
return true;
|
||||||
|
} else if (is_iff(e)) {
|
||||||
|
lhs = arg(e, 1);
|
||||||
|
rhs = arg(e, 2);
|
||||||
|
return true;
|
||||||
|
} else if (is_heq(e)) {
|
||||||
|
lhs = heq_lhs(e);
|
||||||
|
rhs = heq_rhs(e);
|
||||||
|
return true;
|
||||||
|
} else {
|
||||||
|
return false;
|
||||||
|
}
|
||||||
|
}
|
||||||
|
|
||||||
|
expr_pair get_equality_args(expr const & e) {
|
||||||
|
if (is_eq(e)) {
|
||||||
|
return mk_pair(arg(e, 2), arg(e, 3));
|
||||||
|
} else if (is_iff(e)) {
|
||||||
|
return mk_pair(arg(e, 1), arg(e, 2));
|
||||||
|
} else if (is_heq(e)) {
|
||||||
|
return mk_pair(heq_lhs(e), heq_rhs(e));
|
||||||
|
} else {
|
||||||
|
lean_unreachable(); // LCOV_EXCL_LINE
|
||||||
|
}
|
||||||
|
}
|
||||||
|
}
|
|
@ -8,6 +8,7 @@ Author: Leonardo de Moura
|
||||||
#include "library/expr_pair.h"
|
#include "library/expr_pair.h"
|
||||||
|
|
||||||
namespace lean {
|
namespace lean {
|
||||||
bool is_eq_heq(expr const & e);
|
bool is_equality(expr const & e);
|
||||||
expr_pair eq_heq_args(expr const & e);
|
bool is_equality(expr const & e, expr & lhs, expr & rhs);
|
||||||
|
expr_pair get_equality_args(expr const & e);
|
||||||
}
|
}
|
|
@ -9,7 +9,7 @@ Author: Leonardo de Moura
|
||||||
#include "library/fo_unify.h"
|
#include "library/fo_unify.h"
|
||||||
#include "library/expr_pair.h"
|
#include "library/expr_pair.h"
|
||||||
#include "library/kernel_bindings.h"
|
#include "library/kernel_bindings.h"
|
||||||
#include "library/eq_heq.h"
|
#include "library/equality.h"
|
||||||
|
|
||||||
namespace lean {
|
namespace lean {
|
||||||
static void assign(substitution & s, expr const & mvar, expr const & e) {
|
static void assign(substitution & s, expr const & mvar, expr const & e) {
|
||||||
|
@ -36,9 +36,9 @@ optional<substitution> fo_unify(expr e1, expr e2) {
|
||||||
assign(s, e1, e2);
|
assign(s, e1, e2);
|
||||||
} else if (is_metavar_wo_local_context(e2)) {
|
} else if (is_metavar_wo_local_context(e2)) {
|
||||||
assign(s, e2, e1);
|
assign(s, e2, e1);
|
||||||
} else if (is_eq_heq(e1) && is_eq_heq(e2)) {
|
} else if (is_equality(e1) && is_equality(e2)) {
|
||||||
expr_pair p1 = eq_heq_args(e1);
|
expr_pair p1 = get_equality_args(e1);
|
||||||
expr_pair p2 = eq_heq_args(e2);
|
expr_pair p2 = get_equality_args(e2);
|
||||||
todo.emplace_back(p1.second, p2.second);
|
todo.emplace_back(p1.second, p2.second);
|
||||||
todo.emplace_back(p1.first, p2.first);
|
todo.emplace_back(p1.first, p2.first);
|
||||||
} else {
|
} else {
|
||||||
|
|
|
@ -8,7 +8,8 @@ Author: Leonardo de Moura
|
||||||
#include "util/buffer.h"
|
#include "util/buffer.h"
|
||||||
#include "kernel/free_vars.h"
|
#include "kernel/free_vars.h"
|
||||||
#include "kernel/instantiate.h"
|
#include "kernel/instantiate.h"
|
||||||
#include "library/eq_heq.h"
|
#include "kernel/kernel.h"
|
||||||
|
#include "library/equality.h"
|
||||||
#include "library/kernel_bindings.h"
|
#include "library/kernel_bindings.h"
|
||||||
|
|
||||||
namespace lean {
|
namespace lean {
|
||||||
|
@ -183,13 +184,13 @@ class hop_match_fn {
|
||||||
return true;
|
return true;
|
||||||
}
|
}
|
||||||
|
|
||||||
if (is_eq_heq(p) && is_eq_heq(t) && (is_heq(p) || is_heq(t))) {
|
if (is_equality(p) && is_equality(t) && (!is_eq(p) || !is_eq(t))) {
|
||||||
// Remark: if p and t are homogeneous equality, then we handle as an application (in the else branch)
|
// Remark: if p and t are homogeneous equality, then we handle as an application (in the else branch)
|
||||||
// We do that because we can get more information. For example, the pattern
|
// We do that because we can get more information. For example, the pattern
|
||||||
// may be (eq #1 a b).
|
// may be (eq #1 a b).
|
||||||
// This branch ignores the type.
|
// This branch ignores the type.
|
||||||
expr_pair p1 = eq_heq_args(p);
|
expr_pair p1 = get_equality_args(p);
|
||||||
expr_pair p2 = eq_heq_args(t);
|
expr_pair p2 = get_equality_args(t);
|
||||||
return match(p1.first, p2.first, ctx, ctx_size) && match(p1.second, p2.second, ctx, ctx_size);
|
return match(p1.first, p2.first, ctx, ctx_size) && match(p1.second, p2.second, ctx, ctx_size);
|
||||||
} else {
|
} else {
|
||||||
if (p.kind() != t.kind())
|
if (p.kind() != t.kind())
|
||||||
|
|
|
@ -11,6 +11,7 @@ Author: Leonardo de Moura
|
||||||
#include "library/expr_pair.h"
|
#include "library/expr_pair.h"
|
||||||
#include "library/ite.h"
|
#include "library/ite.h"
|
||||||
#include "library/kernel_bindings.h"
|
#include "library/kernel_bindings.h"
|
||||||
|
#include "library/equality.h"
|
||||||
|
|
||||||
namespace lean {
|
namespace lean {
|
||||||
static name g_Hc("Hc"); // auxiliary name for if-then-else
|
static name g_Hc("Hc"); // auxiliary name for if-then-else
|
||||||
|
@ -44,18 +45,18 @@ class to_ceqs_fn {
|
||||||
}
|
}
|
||||||
|
|
||||||
list<expr_pair> apply(expr const & e, expr const & H) {
|
list<expr_pair> apply(expr const & e, expr const & H) {
|
||||||
if (is_eq(e)) {
|
if (is_equality(e)) {
|
||||||
return mk_singleton(e, H);
|
return mk_singleton(e, H);
|
||||||
} else if (is_neq(e)) {
|
} else if (is_neq(e)) {
|
||||||
expr T = arg(e, 1);
|
expr T = arg(e, 1);
|
||||||
expr lhs = arg(e, 2);
|
expr lhs = arg(e, 2);
|
||||||
expr rhs = arg(e, 3);
|
expr rhs = arg(e, 3);
|
||||||
expr new_e = mk_eq(Bool, mk_eq(T, lhs, rhs), False);
|
expr new_e = mk_iff(mk_eq(T, lhs, rhs), False);
|
||||||
expr new_H = mk_neq_elim_th(T, lhs, rhs, H);
|
expr new_H = mk_neq_elim_th(T, lhs, rhs, H);
|
||||||
return mk_singleton(new_e, new_H);
|
return mk_singleton(new_e, new_H);
|
||||||
} else if (is_not(e)) {
|
} else if (is_not(e)) {
|
||||||
expr a = arg(e, 1);
|
expr a = arg(e, 1);
|
||||||
expr new_e = mk_eq(Bool, a, False);
|
expr new_e = mk_iff(a, False);
|
||||||
expr new_H = mk_eqf_intro_th(a, H);
|
expr new_H = mk_eqf_intro_th(a, H);
|
||||||
return mk_singleton(new_e, new_H);
|
return mk_singleton(new_e, new_H);
|
||||||
} else if (is_and(e)) {
|
} else if (is_and(e)) {
|
||||||
|
@ -99,7 +100,7 @@ class to_ceqs_fn {
|
||||||
});
|
});
|
||||||
return append(new_then_ceqs, new_else_ceqs);
|
return append(new_then_ceqs, new_else_ceqs);
|
||||||
} else {
|
} else {
|
||||||
return mk_singleton(mk_eq(Bool, e, True), mk_eqt_intro_th(e, H));
|
return mk_singleton(mk_iff(e, True), mk_eqt_intro_th(e, H));
|
||||||
}
|
}
|
||||||
}
|
}
|
||||||
public:
|
public:
|
||||||
|
@ -146,8 +147,8 @@ bool is_ceq(ro_environment const & env, expr e) {
|
||||||
ctx = extend(ctx, abst_name(e), abst_domain(e));
|
ctx = extend(ctx, abst_name(e), abst_domain(e));
|
||||||
e = abst_body(e);
|
e = abst_body(e);
|
||||||
}
|
}
|
||||||
if (is_eq(e)) {
|
expr lhs, rhs;
|
||||||
expr lhs = arg(e, 2);
|
if (is_equality(e, lhs, rhs)) {
|
||||||
// traverse lhs, and mark all variables that occur there in is_lhs.
|
// traverse lhs, and mark all variables that occur there in is_lhs.
|
||||||
for_each(lhs, visitor_fn);
|
for_each(lhs, visitor_fn);
|
||||||
return std::find(found_args.begin(), found_args.end(), false) == found_args.end();
|
return std::find(found_args.begin(), found_args.end(), false) == found_args.end();
|
||||||
|
@ -211,13 +212,14 @@ bool is_permutation_ceq(expr e) {
|
||||||
e = abst_body(e);
|
e = abst_body(e);
|
||||||
num_args++;
|
num_args++;
|
||||||
}
|
}
|
||||||
if (!is_eq(e))
|
expr lhs, rhs;
|
||||||
return false;
|
if (is_equality(e, lhs, rhs)) {
|
||||||
expr lhs = arg(e, 2);
|
|
||||||
expr rhs = arg(e, 3);
|
|
||||||
buffer<optional<unsigned>> permutation;
|
buffer<optional<unsigned>> permutation;
|
||||||
permutation.resize(num_args);
|
permutation.resize(num_args);
|
||||||
return is_permutation(lhs, rhs, 0, permutation);
|
return is_permutation(lhs, rhs, 0, permutation);
|
||||||
|
} else {
|
||||||
|
return false;
|
||||||
|
}
|
||||||
}
|
}
|
||||||
|
|
||||||
static int to_ceqs(lua_State * L) {
|
static int to_ceqs(lua_State * L) {
|
||||||
|
|
Loading…
Reference in a new issue