refactor(library/logic/connectives): rename theorems

This commit is contained in:
Jeremy Avigad 2014-12-15 15:05:44 -05:00
parent 6903118900
commit 3e9a484851
23 changed files with 195 additions and 190 deletions

View file

@ -537,7 +537,7 @@ We now use =not_intro= and =absurd= to produce a proof term for
import logic import logic
constants a b : Prop constants a b : Prop
check fun (Hab : a → b) (Hnb : ¬ b), check fun (Hab : a → b) (Hnb : ¬ b),
not_intro (fun Ha : a, absurd (Hab Ha) Hnb) not.intro (fun Ha : a, absurd (Hab Ha) Hnb)
#+END_SRC #+END_SRC
@ -743,7 +743,7 @@ if we can derive =B= using an "abstract" witness =w= and a proof term =Hw : B w=
#+BEGIN_SRC lean #+BEGIN_SRC lean
import logic import logic
check @exists_elim check @exists_elim
#+END_SRC #+END_SRC
In the following example, we define =even a= as =∃ b, a = 2*b=, and then we show that the sum In the following example, we define =even a= as =∃ b, a = 2*b=, and then we show that the sum

View file

@ -252,7 +252,7 @@ strong_order_pair.mk strict_order_with_le.le
(assume H : a ≤ b ∧ a ≠ b, (assume H : a ≤ b ∧ a ≠ b,
have H1 : a < b a = b, have H1 : a < b a = b,
from iff.mp !strict_order_with_le.le_iff_lt_or_eq (and.elim_left H), from iff.mp !strict_order_with_le.le_iff_lt_or_eq (and.elim_left H),
show a < b, from or.resolve_left H1 (and.elim_right H))) show a < b, from or_resolve_left H1 (and.elim_right H)))
strict_order_with_le.le_iff_lt_or_eq strict_order_with_le.le_iff_lt_or_eq

View file

@ -321,13 +321,13 @@ section
theorem mul.cancel_right {a b c : A} (Ha : a ≠ 0) (H : b * a = c * a) : b = c := theorem mul.cancel_right {a b c : A} (Ha : a ≠ 0) (H : b * a = c * a) : b = c :=
have H1 : b * a - c * a = 0, from iff.mp !eq_iff_sub_eq_zero H, have H1 : b * a - c * a = 0, from iff.mp !eq_iff_sub_eq_zero H,
have H2 : (b - c) * a = 0, from eq.trans !mul_sub_right_distrib H1, have H2 : (b - c) * a = 0, from eq.trans !mul_sub_right_distrib H1,
have H3 : b - c = 0, from or.resolve_left (eq_zero_or_eq_zero_of_mul_eq_zero H2) Ha, have H3 : b - c = 0, from or_resolve_left (eq_zero_or_eq_zero_of_mul_eq_zero H2) Ha,
iff.elim_right !eq_iff_sub_eq_zero H3 iff.elim_right !eq_iff_sub_eq_zero H3
theorem mul.cancel_left {a b c : A} (Ha : a ≠ 0) (H : a * b = a * c) : b = c := theorem mul.cancel_left {a b c : A} (Ha : a ≠ 0) (H : a * b = a * c) : b = c :=
have H1 : a * b - a * c = 0, from iff.mp !eq_iff_sub_eq_zero H, have H1 : a * b - a * c = 0, from iff.mp !eq_iff_sub_eq_zero H,
have H2 : a * (b - c) = 0, from eq.trans !mul_sub_left_distrib H1, have H2 : a * (b - c) = 0, from eq.trans !mul_sub_left_distrib H1,
have H3 : b - c = 0, from or.resolve_right (eq_zero_or_eq_zero_of_mul_eq_zero H2) Ha, have H3 : b - c = 0, from or_resolve_right (eq_zero_or_eq_zero_of_mul_eq_zero H2) Ha,
iff.elim_right !eq_iff_sub_eq_zero H3 iff.elim_right !eq_iff_sub_eq_zero H3
-- TODO: do we want the iff versions? -- TODO: do we want the iff versions?

View file

@ -42,7 +42,7 @@ open eq.ops
namespace nat namespace nat
theorem succ_pred_of_pos {n : } (H : n > 0) : succ (pred n) = n := theorem succ_pred_of_pos {n : } (H : n > 0) : succ (pred n) = n :=
(or.resolve_right (zero_or_succ_pred n) (ne.symm (lt_imp_ne H))⁻¹) (or_resolve_right (zero_or_succ_pred n) (ne.symm (lt_imp_ne H))⁻¹)
theorem sub_pos_of_gt {m n : } (H : n > m) : n - m > 0 := theorem sub_pos_of_gt {m n : } (H : n > m) : n - m > 0 :=
have H1 : n = n - m + m, from (add_sub_ge_left (lt_imp_le H))⁻¹, have H1 : n = n - m + m, from (add_sub_ge_left (lt_imp_le H))⁻¹,
@ -866,23 +866,23 @@ have H2 : (nat_abs a) * (nat_abs b) = 0, from
... = (nat_abs 0) : {H} ... = (nat_abs 0) : {H}
... = 0 : nat_abs_of_nat 0, ... = 0 : nat_abs_of_nat 0,
have H3 : (nat_abs a) = 0 (nat_abs b) = 0, from mul.eq_zero H2, have H3 : (nat_abs a) = 0 (nat_abs b) = 0, from mul.eq_zero H2,
or.imp_or H3 or_of_or_of_imp_of_imp H3
(assume H : (nat_abs a) = 0, nat_abs_eq_zero H) (assume H : (nat_abs a) = 0, nat_abs_eq_zero H)
(assume H : (nat_abs b) = 0, nat_abs_eq_zero H) (assume H : (nat_abs b) = 0, nat_abs_eq_zero H)
theorem mul_cancel_left_or {a b c : } (H : a * b = a * c) : a = 0 b = c := theorem mul_cancel_left_or {a b c : } (H : a * b = a * c) : a = 0 b = c :=
have H2 : a * (b - c) = 0, by simp, have H2 : a * (b - c) = 0, by simp,
have H3 : a = 0 b - c = 0, from mul_eq_zero H2, have H3 : a = 0 b - c = 0, from mul_eq_zero H2,
or.imp_or_right H3 (assume H4 : b - c = 0, sub_eq_zero H4) or_of_or_of_imp_right H3 (assume H4 : b - c = 0, sub_eq_zero H4)
theorem mul_cancel_left {a b c : } (H1 : a ≠ 0) (H2 : a * b = a * c) : b = c := theorem mul_cancel_left {a b c : } (H1 : a ≠ 0) (H2 : a * b = a * c) : b = c :=
or.resolve_right (mul_cancel_left_or H2) H1 or_resolve_right (mul_cancel_left_or H2) H1
theorem mul_cancel_right_or {a b c : } (H : b * a = c * a) : a = 0 b = c := theorem mul_cancel_right_or {a b c : } (H : b * a = c * a) : a = 0 b = c :=
mul_cancel_left_or ((H ▸ (mul_comm b a)) ▸ mul_comm c a) mul_cancel_left_or ((H ▸ (mul_comm b a)) ▸ mul_comm c a)
theorem mul_cancel_right {a b c : } (H1 : c ≠ 0) (H2 : a * c = b * c) : a = b := theorem mul_cancel_right {a b c : } (H1 : c ≠ 0) (H2 : a * c = b * c) : a = b :=
or.resolve_right (mul_cancel_right_or H2) H1 or_resolve_right (mul_cancel_right_or H2) H1
theorem mul_ne_zero {a b : } (Ha : a ≠ 0) (Hb : b ≠ 0) : a * b ≠ 0 := theorem mul_ne_zero {a b : } (Ha : a ≠ 0) (Hb : b ≠ 0) : a * b ≠ 0 :=
(assume H : a * b = 0, (assume H : a * b = 0,

View file

@ -18,7 +18,7 @@ open int eq.ops
namespace int namespace int
theorem nonneg_elim {a : } : nonneg a → ∃n : , a = n := theorem nonneg_elim {a : } : nonneg a → ∃n : , a = n :=
cases_on a (take n H, exists_intro n rfl) (take n' H, false_elim H) cases_on a (take n H, exists_intro n rfl) (take n' H, false.elim H)
theorem le_intro {a b : } {n : } (H : a + n = b) : a ≤ b := theorem le_intro {a b : } {n : } (H : a + n = b) : a ≤ b :=
have H1 : b - a = n, from add_imp_sub_right (!add_comm ▸ H), have H1 : b - a = n, from add_imp_sub_right (!add_comm ▸ H),
@ -262,7 +262,7 @@ theorem le_imp_lt_or_eq {a b : } (H : a ≤ b) : a < b a = b :=
le_imp_succ_le_or_eq H le_imp_succ_le_or_eq H
theorem le_ne_imp_lt {a b : } (H1 : a ≤ b) (H2 : a ≠ b) : a < b := theorem le_ne_imp_lt {a b : } (H1 : a ≤ b) (H2 : a ≠ b) : a < b :=
or.resolve_left (le_imp_lt_or_eq H1) H2 or_resolve_left (le_imp_lt_or_eq H1) H2
theorem le_imp_lt_succ {a b : } (H : a ≤ b) : a < b + 1 := theorem le_imp_lt_succ {a b : } (H : a ≤ b) : a < b + 1 :=
add_le_right H 1 add_le_right H 1
@ -379,26 +379,26 @@ by_cases a
or.inl (neg_le_pos n m)) or.inl (neg_le_pos n m))
(take m : , (take m : ,
show -n ≤ -succ m -n > -succ m, from show -n ≤ -succ m -n > -succ m, from
or.imp_or le_or_gt or_of_or_of_imp_of_imp le_or_gt
(assume H : succ m ≤ n, (assume H : succ m ≤ n,
le_neg (iff.elim_left (iff.symm (le_of_nat (succ m) n)) H)) le_neg (iff.elim_left (iff.symm (le_of_nat (succ m) n)) H))
(assume H : succ m > n, (assume H : succ m > n,
lt_neg (iff.elim_left (iff.symm (lt_of_nat n (succ m))) H)))) lt_neg (iff.elim_left (iff.symm (lt_of_nat n (succ m))) H))))
theorem trichotomy_alt (a b : ) : (a < b a = b) a > b := theorem trichotomy_alt (a b : ) : (a < b a = b) a > b :=
or.imp_or_left (le_or_gt a b) (assume H : a ≤ b, le_imp_lt_or_eq H) or_of_or_of_imp_left (le_or_gt a b) (assume H : a ≤ b, le_imp_lt_or_eq H)
theorem trichotomy (a b : ) : a < b a = b a > b := theorem trichotomy (a b : ) : a < b a = b a > b :=
iff.elim_left or.assoc (trichotomy_alt a b) iff.elim_left or.assoc (trichotomy_alt a b)
theorem le_total (a b : ) : a ≤ b b ≤ a := theorem le_total (a b : ) : a ≤ b b ≤ a :=
or.imp_or_right (le_or_gt a b) (assume H : b < a, lt_imp_le H) or_of_or_of_imp_right (le_or_gt a b) (assume H : b < a, lt_imp_le H)
theorem not_lt_imp_le {a b : } (H : ¬ a < b) : b ≤ a := theorem not_lt_imp_le {a b : } (H : ¬ a < b) : b ≤ a :=
or.resolve_left (le_or_gt b a) H or_resolve_left (le_or_gt b a) H
theorem not_le_imp_lt {a b : } (H : ¬ a ≤ b) : b < a := theorem not_le_imp_lt {a b : } (H : ¬ a ≤ b) : b < a :=
or.resolve_right (le_or_gt a b) H or_resolve_right (le_or_gt a b) H
-- (non)positivity and (non)negativity -- (non)positivity and (non)negativity
-- ------------------------------------- -- -------------------------------------
@ -448,7 +448,7 @@ calc
... = -a : (neg_move (Hn⁻¹))⁻¹ ... = -a : (neg_move (Hn⁻¹))⁻¹
theorem nat_abs_cases (a : ) : a = (nat_abs a) a = - (nat_abs a) := theorem nat_abs_cases (a : ) : a = (nat_abs a) a = - (nat_abs a) :=
or.imp_or (le_total 0 a) or_of_or_of_imp_of_imp (le_total 0 a)
(assume H : a ≥ 0, (nat_abs_nonneg_eq H)⁻¹) (assume H : a ≥ 0, (nat_abs_nonneg_eq H)⁻¹)
(assume H : a ≤ 0, (neg_move ((nat_abs_negative H)⁻¹))⁻¹) (assume H : a ≤ 0, (neg_move ((nat_abs_negative H)⁻¹))⁻¹)
@ -564,7 +564,7 @@ have H2 : (nat_abs a) * (nat_abs b) = 1, from
... = (nat_abs 1) : {H} ... = (nat_abs 1) : {H}
... = 1 : nat_abs_of_nat 1, ... = 1 : nat_abs_of_nat 1,
have H3 : (nat_abs a) = 1, from mul_eq_one_left H2, have H3 : (nat_abs a) = 1, from mul_eq_one_left H2,
or.imp_or (nat_abs_cases a) or_of_or_of_imp_of_imp (nat_abs_cases a)
(assume H4 : a = (nat_abs a), H3 ▸ H4) (assume H4 : a = (nat_abs a), H3 ▸ H4)
(assume H4 : a = - (nat_abs a), H3 ▸ H4) (assume H4 : a = - (nat_abs a), H3 ▸ H4)

View file

@ -146,12 +146,12 @@ induction_on s or.inr
assume IH : x ∈ s ++ t → x ∈ s x ∈ t, assume IH : x ∈ s ++ t → x ∈ s x ∈ t,
assume H1 : x ∈ y::s ++ t, assume H1 : x ∈ y::s ++ t,
have H2 : x = y x ∈ s ++ t, from H1, have H2 : x = y x ∈ s ++ t, from H1,
have H3 : x = y x ∈ s x ∈ t, from or.imp_or_right H2 IH, have H3 : x = y x ∈ s x ∈ t, from or_of_or_of_imp_right H2 IH,
iff.elim_right or.assoc H3) iff.elim_right or.assoc H3)
theorem mem.or_imp_concat {x : T} {s t : list T} : x ∈ s x ∈ t → x ∈ s ++ t := theorem mem.or_imp_concat {x : T} {s t : list T} : x ∈ s x ∈ t → x ∈ s ++ t :=
induction_on s induction_on s
(take H, or.elim H false_elim (assume H, H)) (take H, or.elim H false.elim (assume H, H))
(take y s, (take y s,
assume IH : x ∈ s x ∈ t → x ∈ s ++ t, assume IH : x ∈ s x ∈ t → x ∈ s ++ t,
assume H : x ∈ y::s x ∈ t, assume H : x ∈ y::s x ∈ t,
@ -167,7 +167,7 @@ iff.intro mem.concat_imp_or mem.or_imp_concat
theorem mem.split {x : T} {l : list T} : x ∈ l → ∃s t : list T, l = s ++ (x::t) := theorem mem.split {x : T} {l : list T} : x ∈ l → ∃s t : list T, l = s ++ (x::t) :=
induction_on l induction_on l
(take H : x ∈ nil, false_elim (iff.elim_left !mem.nil H)) (take H : x ∈ nil, false.elim (iff.elim_left !mem.nil H))
(take y l, (take y l,
assume IH : x ∈ l → ∃s t : list T, l = s ++ (x::t), assume IH : x ∈ l → ∃s t : list T, l = s ++ (x::t),
assume H : x ∈ y::l, assume H : x ∈ y::l,
@ -226,7 +226,7 @@ rec_on l
assume iH : ¬x ∈ l → find x l = length l, assume iH : ¬x ∈ l → find x l = length l,
assume P₁ : ¬x ∈ y::l, assume P₁ : ¬x ∈ y::l,
have P₂ : ¬(x = y x ∈ l), from iff.elim_right (iff.flip_sign !mem.cons) P₁, have P₂ : ¬(x = y x ∈ l), from iff.elim_right (iff.flip_sign !mem.cons) P₁,
have P₃ : ¬x = y ∧ ¬x ∈ l, from (iff.elim_left not_or P₂), have P₃ : ¬x = y ∧ ¬x ∈ l, from (iff.elim_left not_or_iff_not_and_not P₂),
calc calc
find x (y::l) = if x = y then 0 else succ (find x l) : !find.cons find x (y::l) = if x = y then 0 else succ (find x l) : !find.cons
... = succ (find x l) : if_neg (and.elim_left P₃) ... = succ (find x l) : if_neg (and.elim_left P₃)

View file

@ -63,7 +63,7 @@ induction_on n
(show succ m = succ (pred (succ m)), from congr_arg succ !pred.succ⁻¹)) (show succ m = succ (pred (succ m)), from congr_arg succ !pred.succ⁻¹))
theorem zero_or_exists_succ (n : ) : n = 0 ∃k, n = succ k := theorem zero_or_exists_succ (n : ) : n = 0 ∃k, n = succ k :=
or.imp_or (zero_or_succ_pred n) (assume H, H) or_of_or_of_imp_of_imp (zero_or_succ_pred n) (assume H, H)
(assume H : n = succ (pred n), exists_intro (pred n) H) (assume H : n = succ (pred n), exists_intro (pred n) H)
theorem case {P : → Prop} (n : ) (H1: P 0) (H2 : ∀m, P (succ m)) : P n := theorem case {P : → Prop} (n : ) (H1: P 0) (H2 : ∀m, P (succ m)) : P n :=

View file

@ -34,10 +34,10 @@ congr_fun (fix_eq div.F x) y
notation a div b := divide a b notation a div b := divide a b
theorem div_zero (a : ) : a div 0 = 0 := theorem div_zero (a : ) : a div 0 = 0 :=
divide_def a 0 ⬝ if_neg (!and.not_left (lt.irrefl 0)) divide_def a 0 ⬝ if_neg (!not_and_of_not_left (lt.irrefl 0))
theorem div_less {a b : } (h : a < b) : a div b = 0 := theorem div_less {a b : } (h : a < b) : a div b = 0 :=
divide_def a b ⬝ if_neg (!and.not_right (lt_imp_not_ge h)) divide_def a b ⬝ if_neg (!not_and_of_not_right (lt_imp_not_ge h))
theorem zero_div (b : ) : 0 div b = 0 := theorem zero_div (b : ) : 0 div b = 0 :=
divide_def 0 b ⬝ if_neg (λ h, and.rec_on h (λ l r, absurd (lt.of_lt_of_le l r) (lt.irrefl 0))) divide_def 0 b ⬝ if_neg (λ h, and.rec_on h (λ l r, absurd (lt.of_lt_of_le l r) (lt.irrefl 0)))
@ -74,10 +74,10 @@ theorem modulo_def (x y : nat) : modulo x y = if 0 < y ∧ y ≤ x then modulo (
congr_fun (fix_eq mod.F x) y congr_fun (fix_eq mod.F x) y
theorem mod_zero (a : ) : a mod 0 = a := theorem mod_zero (a : ) : a mod 0 = a :=
modulo_def a 0 ⬝ if_neg (!and.not_left (lt.irrefl 0)) modulo_def a 0 ⬝ if_neg (!not_and_of_not_left (lt.irrefl 0))
theorem mod_less {a b : } (h : a < b) : a mod b = a := theorem mod_less {a b : } (h : a < b) : a mod b = a :=
modulo_def a b ⬝ if_neg (!and.not_right (lt_imp_not_ge h)) modulo_def a b ⬝ if_neg (!not_and_of_not_right (lt_imp_not_ge h))
theorem zero_mod (b : ) : 0 mod b = 0 := theorem zero_mod (b : ) : 0 mod b = 0 :=
modulo_def 0 b ⬝ if_neg (λ h, and.rec_on h (λ l r, absurd (lt.of_lt_of_le l r) (lt.irrefl 0))) modulo_def 0 b ⬝ if_neg (λ h, and.rec_on h (λ l r, absurd (lt.of_lt_of_le l r) (lt.irrefl 0)))

View file

@ -152,10 +152,10 @@ discriminate
or.inl Hlt) or.inl Hlt)
theorem le_ne_imp_succ_le {n m : } (H1 : n ≤ m) (H2 : n ≠ m) : succ n ≤ m := theorem le_ne_imp_succ_le {n m : } (H1 : n ≤ m) (H2 : n ≠ m) : succ n ≤ m :=
or.resolve_left (le_imp_succ_le_or_eq H1) H2 or_resolve_left (le_imp_succ_le_or_eq H1) H2
theorem le_succ_imp_le_or_eq {n m : } (H : n ≤ succ m) : n ≤ m n = succ m := theorem le_succ_imp_le_or_eq {n m : } (H : n ≤ succ m) : n ≤ m n = succ m :=
or.imp_or_left (le_imp_succ_le_or_eq H) or_of_or_of_imp_left (le_imp_succ_le_or_eq H)
(take H2 : succ n ≤ succ m, show n ≤ m, from succ_le_cancel H2) (take H2 : succ n ≤ succ m, show n ≤ m, from succ_le_cancel H2)
theorem succ_le_imp_le_and_ne {n m : } (H : succ n ≤ m) : n ≤ m ∧ n ≠ m := theorem succ_le_imp_le_and_ne {n m : } (H : succ n ≤ m) : n ≤ m ∧ n ≠ m :=
@ -210,7 +210,7 @@ discriminate
have H3 : k ≤ m, from H2 ▸ H, have H3 : k ≤ m, from H2 ▸ H,
have H4 : succ k ≤ m k = m, from le_imp_succ_le_or_eq H3, have H4 : succ k ≤ m k = m, from le_imp_succ_le_or_eq H3,
show n ≤ m n = succ m, from show n ≤ m n = succ m, from
or.imp_or H4 or_of_or_of_imp_of_imp H4
(take H5 : succ k ≤ m, show n ≤ m, from Hn⁻¹ ▸ H5) (take H5 : succ k ≤ m, show n ≤ m, from Hn⁻¹ ▸ H5)
(take H5 : k = m, show n = succ m, from H5 ▸ Hn)) (take H5 : k = m, show n = succ m, from H5 ▸ Hn))
@ -275,7 +275,7 @@ theorem le_imp_lt_or_eq {n m : } (H : n ≤ m) : n < m n = m :=
or.swap (eq_or_lt_of_le H) or.swap (eq_or_lt_of_le H)
theorem le_ne_imp_lt {n m : } (H1 : n ≤ m) (H2 : n ≠ m) : n < m := theorem le_ne_imp_lt {n m : } (H1 : n ≤ m) (H2 : n ≠ m) : n < m :=
or.resolve_left (le_imp_lt_or_eq H1) H2 or_resolve_left (le_imp_lt_or_eq H1) H2
theorem lt_succ_imp_le {n m : } (H : n < succ m) : n ≤ m := theorem lt_succ_imp_le {n m : } (H : n < succ m) : n ≤ m :=
succ_le_cancel (succ_le_of_lt H) succ_le_cancel (succ_le_of_lt H)
@ -347,13 +347,13 @@ theorem trichotomy (n m : ) : n < m n = m n > m :=
lt.trichotomy n m lt.trichotomy n m
theorem le_total (n m : ) : n ≤ m m ≤ n := theorem le_total (n m : ) : n ≤ m m ≤ n :=
or.imp_or_right le_or_gt (assume H : m < n, lt_imp_le H) or_of_or_of_imp_right le_or_gt (assume H : m < n, lt_imp_le H)
theorem not_lt_imp_ge {n m : } (H : ¬ n < m) : n ≥ m := theorem not_lt_imp_ge {n m : } (H : ¬ n < m) : n ≥ m :=
or.resolve_left le_or_gt H or_resolve_left le_or_gt H
theorem not_le_imp_gt {n m : } (H : ¬ n ≤ m) : n > m := theorem not_le_imp_gt {n m : } (H : ¬ n ≤ m) : n > m :=
or.resolve_right le_or_gt H or_resolve_right le_or_gt H
-- ### misc -- ### misc
@ -396,7 +396,7 @@ theorem case_zero_pos {P : → Prop} (y : ) (H0 : P 0) (H1 : ∀ {y : nat
case y H0 (take y, H1 !succ_pos) case y H0 (take y, H1 !succ_pos)
theorem zero_or_pos {n : } : n = 0 n > 0 := theorem zero_or_pos {n : } : n = 0 n > 0 :=
or.imp_or_left or_of_or_of_imp_left
(or.swap (le_imp_lt_or_eq !zero_le)) (or.swap (le_imp_lt_or_eq !zero_le))
(take H : 0 = n, H⁻¹) (take H : 0 = n, H⁻¹)
@ -493,7 +493,7 @@ have H5 : k ≤ m, from mul_le_cancel_left Hn H3,
le_antisym H4 H5 le_antisym H4 H5
theorem mul_cancel_left_or {n m k : } (H : n * m = n * k) : n = 0 m = k := theorem mul_cancel_left_or {n m k : } (H : n * m = n * k) : n = 0 m = k :=
or.imp_or_right zero_or_pos or_of_or_of_imp_right zero_or_pos
(assume Hn : n > 0, mul_cancel_left Hn H) (assume Hn : n > 0, mul_cancel_left Hn H)
theorem mul_cancel_right {n m k : } (Hm : m > 0) (H : n * m = k * m) : n = k := theorem mul_cancel_right {n m k : } (Hm : m > 0) (H : n * m = k * m) : n = k :=

View file

@ -36,9 +36,9 @@ theorem propext {a b : Prop} (Hab : a → b) (Hba : b → a) : a = b :=
or.elim (prop_complete a) or.elim (prop_complete a)
(assume Hat, or.elim (prop_complete b) (assume Hat, or.elim (prop_complete b)
(assume Hbt, Hat ⬝ Hbt⁻¹) (assume Hbt, Hat ⬝ Hbt⁻¹)
(assume Hbf, false_elim (Hbf ▸ (Hab (of_eq_true Hat))))) (assume Hbf, false.elim (Hbf ▸ (Hab (of_eq_true Hat)))))
(assume Haf, or.elim (prop_complete b) (assume Haf, or.elim (prop_complete b)
(assume Hbt, false_elim (Haf ▸ (Hba (of_eq_true Hbt)))) (assume Hbt, false.elim (Haf ▸ (Hba (of_eq_true Hbt))))
(assume Hbf, Haf ⬝ Hbf⁻¹)) (assume Hbf, Haf ⬝ Hbf⁻¹))
theorem eq.of_iff {a b : Prop} (H : a ↔ b) : a = b := theorem eq.of_iff {a b : Prop} (H : a ↔ b) : a = b :=

View file

@ -1,155 +1,153 @@
-- Copyright (c) 2014 Microsoft Corporation. All rights reserved. /-
-- Released under Apache 2.0 license as described in the file LICENSE. Copyright (c) 2014 Microsoft Corporation. All rights reserved.
-- Authors: Jeremy Avigad, Leonardo de Moura Released under Apache 2.0 license as described in the file LICENSE.
definition imp (a b : Prop) : Prop := a → b Module: logic.connectives
Authors: Jeremy Avigad, Leonardo de Moura
The propositional connectives. See also init.datatypes.
-/
variables {a b c d : Prop} variables {a b c d : Prop}
/- implies -/
definition imp (a b : Prop) : Prop := a → b
theorem mt (H1 : a → b) (H2 : ¬b) : ¬a := theorem mt (H1 : a → b) (H2 : ¬b) : ¬a :=
assume Ha : a, absurd (H1 Ha) H2 assume Ha : a, absurd (H1 Ha) H2
-- make c explicit and rename to false.elim /- false -/
theorem false_elim {c : Prop} (H : false) : c :=
theorem false.elim {c : Prop} (H : false) : c :=
false.rec c H false.rec c H
-- not /- not -/
-- ---
theorem not_elim (H1 : ¬a) (H2 : a) : false := H1 H2 theorem not.elim (H1 : ¬a) (H2 : a) : false := H1 H2
theorem not_intro (H : a → false) : ¬a := H theorem not.intro (H : a → false) : ¬a := H
theorem not_not_intro (Ha : a) : ¬¬a := theorem not_not_intro (Ha : a) : ¬¬a :=
assume Hna : ¬a, absurd Ha Hna assume Hna : ¬a, absurd Ha Hna
theorem not_implies_left (H : ¬(a → b)) : ¬¬a := theorem not_not_of_not_implies (H : ¬(a → b)) : ¬¬a :=
assume Hna : ¬a, absurd (assume Ha : a, absurd Ha Hna) H assume Hna : ¬a, absurd (assume Ha : a, absurd Ha Hna) H
theorem not_implies_right (H : ¬(a → b)) : ¬b := theorem not_of_not_implies (H : ¬(a → b)) : ¬b :=
assume Hb : b, absurd (assume Ha : a, Hb) H assume Hb : b, absurd (assume Ha : a, Hb) H
theorem not_not_em : ¬¬(a ¬a) := theorem not_not_em : ¬¬(a ¬a) :=
assume not_em : ¬(a ¬a), assume not_em : ¬(a ¬a),
have Hnp : ¬a, from have Hnp : ¬a, from
assume Hp : a, absurd (or.inl Hp) not_em, assume Hp : a, absurd (or.inl Hp) not_em,
absurd (or.inr Hnp) not_em absurd (or.inr Hnp) not_em
-- and /- and -/
-- ---
namespace and definition not_and_of_not_left (b : Prop) (Hna : ¬a) : ¬(a ∧ b) :=
definition not_left (b : Prop) (Hna : ¬a) : ¬(a ∧ b) := assume H : a ∧ b, absurd (and.elim_left H) Hna
assume H : a ∧ b, absurd (elim_left H) Hna
definition not_right (a : Prop) {b : Prop} (Hnb : ¬b) : ¬(a ∧ b) := definition not_and_of_not_right (a : Prop) {b : Prop} (Hnb : ¬b) : ¬(a ∧ b) :=
assume H : a ∧ b, absurd (elim_right H) Hnb assume H : a ∧ b, absurd (and.elim_right H) Hnb
theorem swap (H : a ∧ b) : b ∧ a := theorem and.swap (H : a ∧ b) : b ∧ a :=
intro (elim_right H) (elim_left H) and.intro (and.elim_right H) (and.elim_left H)
theorem imp_and (H₁ : a ∧ b) (H₂ : a → c) (H₃ : b → d) : c ∧ d := theorem and_of_and_of_imp_of_imp (H₁ : a ∧ b) (H₂ : a → c) (H₃ : b → d) : c ∧ d :=
elim H₁ (assume Ha : a, assume Hb : b, intro (H₂ Ha) (H₃ Hb)) and.elim H₁ (assume Ha : a, assume Hb : b, and.intro (H₂ Ha) (H₃ Hb))
theorem imp_left (H₁ : a ∧ c) (H : a → b) : b ∧ c := theorem and_of_and_of_imp_left (H₁ : a ∧ c) (H : a → b) : b ∧ c :=
elim H₁ (assume Ha : a, assume Hc : c, intro (H Ha) Hc) and.elim H₁ (assume Ha : a, assume Hc : c, and.intro (H Ha) Hc)
theorem imp_right (H₁ : c ∧ a) (H : a → b) : c ∧ b := theorem and_of_and_of_imp_right (H₁ : c ∧ a) (H : a → b) : c ∧ b :=
elim H₁ (assume Hc : c, assume Ha : a, intro Hc (H Ha)) and.elim H₁ (assume Hc : c, assume Ha : a, and.intro Hc (H Ha))
theorem comm : a ∧ b ↔ b ∧ a := theorem and.comm : a ∧ b ↔ b ∧ a :=
iff.intro (λH, swap H) (λH, swap H) iff.intro (λH, and.swap H) (λH, and.swap H)
theorem assoc : (a ∧ b) ∧ c ↔ a ∧ (b ∧ c) := theorem and.assoc : (a ∧ b) ∧ c ↔ a ∧ (b ∧ c) :=
iff.intro iff.intro
(assume H, intro (assume H, and.intro
(elim_left (elim_left H)) (and.elim_left (and.elim_left H))
(intro (elim_right (elim_left H)) (elim_right H))) (and.intro (and.elim_right (and.elim_left H)) (and.elim_right H)))
(assume H, intro (assume H, and.intro
(intro (elim_left H) (elim_left (elim_right H))) (and.intro (and.elim_left H) (and.elim_left (and.elim_right H)))
(elim_right (elim_right H))) (and.elim_right (and.elim_right H)))
end and
-- or /- or -/
-- --
namespace or definition not_or (Hna : ¬a) (Hnb : ¬b) : ¬(a b) :=
definition not_intro (Hna : ¬a) (Hnb : ¬b) : ¬(a b) := assume H : a b, or.rec_on H
assume H : a b, or.rec_on H (assume Ha, absurd Ha Hna)
(assume Ha, absurd Ha Hna) (assume Hb, absurd Hb Hnb)
(assume Hb, absurd Hb Hnb)
theorem imp_or (H₁ : a b) (H₂ : a → c) (H₃ : b → d) : c d := theorem or_of_or_of_imp_of_imp (H₁ : a b) (H₂ : a → c) (H₃ : b → d) : c d :=
elim H₁ or.elim H₁
(assume Ha : a, inl (H₂ Ha)) (assume Ha : a, or.inl (H₂ Ha))
(assume Hb : b, inr (H₃ Hb)) (assume Hb : b, or.inr (H₃ Hb))
theorem imp_or_left (H₁ : a c) (H : a → b) : b c := theorem or_of_or_of_imp_left (H₁ : a c) (H : a → b) : b c :=
elim H₁ or.elim H₁
(assume H₂ : a, inl (H H₂)) (assume H₂ : a, or.inl (H H₂))
(assume H₂ : c, inr H₂) (assume H₂ : c, or.inr H₂)
theorem imp_or_right (H₁ : c a) (H : a → b) : c b := theorem or_of_or_of_imp_right (H₁ : c a) (H : a → b) : c b :=
elim H₁ or.elim H₁
(assume H₂ : c, inl H₂) (assume H₂ : c, or.inl H₂)
(assume H₂ : a, inr (H H₂)) (assume H₂ : a, or.inr (H H₂))
theorem elim3 (H : a b c) (Ha : a → d) (Hb : b → d) (Hc : c → d) : d := theorem or.elim3 (H : a b c) (Ha : a → d) (Hb : b → d) (Hc : c → d) : d :=
elim H Ha (assume H₂, elim H₂ Hb Hc) or.elim H Ha (assume H₂, or.elim H₂ Hb Hc)
theorem resolve_right (H₁ : a b) (H₂ : ¬a) : b := theorem or_resolve_right (H₁ : a b) (H₂ : ¬a) : b :=
elim H₁ (assume Ha, absurd Ha H₂) (assume Hb, Hb) or.elim H₁ (assume Ha, absurd Ha H₂) (assume Hb, Hb)
theorem resolve_left (H₁ : a b) (H₂ : ¬b) : a := theorem or_resolve_left (H₁ : a b) (H₂ : ¬b) : a :=
elim H₁ (assume Ha, Ha) (assume Hb, absurd Hb H₂) or.elim H₁ (assume Ha, Ha) (assume Hb, absurd Hb H₂)
theorem swap (H : a b) : b a := theorem or.swap (H : a b) : b a :=
elim H (assume Ha, inr Ha) (assume Hb, inl Hb) or.elim H (assume Ha, or.inr Ha) (assume Hb, or.inl Hb)
theorem comm : a b ↔ b a := theorem or.comm : a b ↔ b a :=
iff.intro (λH, swap H) (λH, swap H) iff.intro (λH, or.swap H) (λH, or.swap H)
theorem assoc : (a b) c ↔ a (b c) := theorem or.assoc : (a b) c ↔ a (b c) :=
iff.intro iff.intro
(assume H, elim H (assume H, or.elim H
(assume H₁, elim H₁ (assume H₁, or.elim H₁
(assume Ha, inl Ha) (assume Ha, or.inl Ha)
(assume Hb, inr (inl Hb))) (assume Hb, or.inr (or.inl Hb)))
(assume Hc, inr (inr Hc))) (assume Hc, or.inr (or.inr Hc)))
(assume H, elim H (assume H, or.elim H
(assume Ha, (inl (inl Ha))) (assume Ha, (or.inl (or.inl Ha)))
(assume H₁, elim H₁ (assume H₁, or.elim H₁
(assume Hb, inl (inr Hb)) (assume Hb, or.inl (or.inr Hb))
(assume Hc, inr Hc))) (assume Hc, or.inr Hc)))
end or
-- iff /- iff -/
-- ---
namespace iff definition iff.def : (a ↔ b) = ((a → b) ∧ (b → a)) :=
definition def : (a ↔ b) = ((a → b) ∧ (b → a)) := !eq.refl
!eq.refl
end iff /- exists_unique -/
-- exists_unique
-- -------------
definition exists_unique {A : Type} (p : A → Prop) := definition exists_unique {A : Type} (p : A → Prop) :=
∃x, p x ∧ ∀y, p y → y = x ∃x, p x ∧ ∀y, p y → y = x
notation `∃!` binders `,` r:(scoped P, exists_unique P) := r notation `∃!` binders `,` r:(scoped P, exists_unique P) := r
theorem exists_unique_intro {A : Type} {p : A → Prop} (w : A) (H1 : p w) (H2 : ∀y, p y → y = w) : ∃!x, p x := theorem exists_unique.intro {A : Type} {p : A → Prop} (w : A) (H1 : p w) (H2 : ∀y, p y → y = w) :
∃!x, p x :=
exists_intro w (and.intro H1 H2) exists_intro w (and.intro H1 H2)
theorem exists_unique_elim {A : Type} {p : A → Prop} {b : Prop} theorem exists_unique.elim {A : Type} {p : A → Prop} {b : Prop}
(H2 : ∃!x, p x) (H1 : ∀x, p x → (∀y, p y → y = x) → b) : b := (H2 : ∃!x, p x) (H1 : ∀x, p x → (∀y, p y → y = x) → b) : b :=
obtain w Hw, from H2, obtain w Hw, from H2,
H1 w (and.elim_left Hw) (and.elim_right Hw) H1 w (and.elim_left Hw) (and.elim_right Hw)
-- if-then-else /- if-then-else -/
-- ------------
section section
open eq.ops open eq.ops
@ -161,8 +159,8 @@ section
definition if_false (t e : A) : (if false then t else e) = e := definition if_false (t e : A) : (if false then t else e) = e :=
if_neg not_false if_neg not_false
theorem if_cond_congr [H₁ : decidable c₁] [H₂ : decidable c₂] (Heq : c₁ ↔ c₂) (t e : A) theorem if_congr_cond [H₁ : decidable c₁] [H₂ : decidable c₂] (Heq : c₁ ↔ c₂) (t e : A) :
: (if c₁ then t else e) = (if c₂ then t else e) := (if c₁ then t else e) = (if c₂ then t else e) :=
decidable.rec_on H₁ decidable.rec_on H₁
(λ Hc₁ : c₁, decidable.rec_on H₂ (λ Hc₁ : c₁, decidable.rec_on H₂
(λ Hc₂ : c₂, if_pos Hc₁ ⬝ (if_pos Hc₂)⁻¹) (λ Hc₂ : c₂, if_pos Hc₁ ⬝ (if_pos Hc₂)⁻¹)
@ -172,12 +170,13 @@ section
(λ Hnc₂ : ¬c₂, if_neg Hnc₁ ⬝ (if_neg Hnc₂)⁻¹)) (λ Hnc₂ : ¬c₂, if_neg Hnc₁ ⬝ (if_neg Hnc₂)⁻¹))
theorem if_congr_aux [H₁ : decidable c₁] [H₂ : decidable c₂] {t₁ t₂ e₁ e₂ : A} theorem if_congr_aux [H₁ : decidable c₁] [H₂ : decidable c₂] {t₁ t₂ e₁ e₂ : A}
(Hc : c₁ ↔ c₂) (Ht : t₁ = t₂) (He : e₁ = e₂) : (Hc : c₁ ↔ c₂) (Ht : t₁ = t₂) (He : e₁ = e₂) :
(if c₁ then t₁ else e₁) = (if c₂ then t₂ else e₂) := (if c₁ then t₁ else e₁) = (if c₂ then t₂ else e₂) :=
Ht ▸ He ▸ (if_cond_congr Hc t₁ e₁) Ht ▸ He ▸ (if_congr_cond Hc t₁ e₁)
theorem if_congr [H₁ : decidable c₁] {t₁ t₂ e₁ e₂ : A} (Hc : c₁ ↔ c₂) (Ht : t₁ = t₂) (He : e₁ = e₂) : theorem if_congr [H₁ : decidable c₁] {t₁ t₂ e₁ e₂ : A} (Hc : c₁ ↔ c₂) (Ht : t₁ = t₂)
(if c₁ then t₁ else e₁) = (@ite c₂ (decidable.decidable_iff_equiv H₁ Hc) A t₂ e₂) := (He : e₁ = e₂) :
(if c₁ then t₁ else e₁) = (@ite c₂ (decidable.decidable_iff_equiv H₁ Hc) A t₂ e₂) :=
have H2 [visible] : decidable c₂, from (decidable.decidable_iff_equiv H₁ Hc), have H2 [visible] : decidable c₂, from (decidable.decidable_iff_equiv H₁ Hc),
if_congr_aux Hc Ht He if_congr_aux Hc Ht He
end end

View file

@ -185,7 +185,7 @@ assume Hem H1,
have Hx : ∀x, P x, from have Hx : ∀x, P x, from
take x, take x,
have H1 : P x C, from H1 x, have H1 : P x C, from H1 x,
or.resolve_left H1 Hnc, or_resolve_left H1 Hnc,
or.inl Hx) or.inl Hx)
theorem thm23a : (∃x, P x) ∧ C → (∃x, P x ∧ C) := theorem thm23a : (∃x, P x) ∧ C → (∃x, P x ∧ C) :=

View file

@ -47,12 +47,13 @@ theorem not_not_elim {a : Prop} [D : decidable a] (H : ¬¬a) : a :=
iff.mp not_not_iff H iff.mp not_not_iff H
theorem not_true_iff_false : (¬true) ↔ false := theorem not_true_iff_false : (¬true) ↔ false :=
iff.intro (assume H, H trivial) false_elim iff.intro (assume H, H trivial) false.elim
theorem not_false_iff_true : (¬false) ↔ true := theorem not_false_iff_true : (¬false) ↔ true :=
iff.intro (assume H, trivial) (assume H H', H') iff.intro (assume H, trivial) (assume H H', H')
theorem not_or {a b : Prop} [Da : decidable a] [Db : decidable b] : (¬(a b)) ↔ (¬a ∧ ¬b) := theorem not_or_iff_not_and_not {a b : Prop} [Da : decidable a] [Db : decidable b] :
(¬(a b)) ↔ (¬a ∧ ¬b) :=
iff.intro iff.intro
(assume H, or.elim (em a) (assume H, or.elim (em a)
(assume Ha, absurd (or.inl Ha) H) (assume Ha, absurd (or.inl Ha) H)
@ -82,16 +83,16 @@ iff.intro
(assume Ha : a, or.inr (H Ha)) (assume Ha : a, or.inr (H Ha))
(assume Hna : ¬a, or.inl Hna))) (assume Hna : ¬a, or.inl Hna)))
(assume (H : ¬a b) (Ha : a), (assume (H : ¬a b) (Ha : a),
or.resolve_right H (not_not_iff⁻¹ ▸ Ha)) or_resolve_right H (not_not_iff⁻¹ ▸ Ha))
theorem not_implies {a b : Prop} [Da : decidable a] [Db : decidable b] : (¬(a → b)) ↔ (a ∧ ¬b) := theorem not_implies {a b : Prop} [Da : decidable a] [Db : decidable b] : (¬(a → b)) ↔ (a ∧ ¬b) :=
calc (¬(a → b)) ↔ (¬(¬a b)) : {imp_or} calc (¬(a → b)) ↔ (¬(¬a b)) : {imp_or}
... ↔ (¬¬a ∧ ¬b) : not_or ... ↔ (¬¬a ∧ ¬b) : not_or_iff_not_and_not
... ↔ (a ∧ ¬b) : {not_not_iff} ... ↔ (a ∧ ¬b) : {not_not_iff}
theorem peirce {a b : Prop} [D : decidable a] : ((a → b) → a) → a := theorem peirce {a b : Prop} [D : decidable a] : ((a → b) → a) → a :=
assume H, by_contradiction (assume Hna : ¬a, assume H, by_contradiction (assume Hna : ¬a,
have Hnna : ¬¬a, from not_implies_left (mt H Hna), have Hnna : ¬¬a, from not_not_of_not_implies (mt H Hna),
absurd (not_not_elim Hnna) Hna) absurd (not_not_elim Hnna) Hna)
theorem not_exists_forall {A : Type} {P : A → Prop} [D : ∀x, decidable (P x)] theorem not_exists_forall {A : Type} {P : A → Prop} [D : ∀x, decidable (P x)]
@ -116,12 +117,12 @@ iff.intro
theorem iff_false_intro {a : Prop} (H : ¬a) : a ↔ false := theorem iff_false_intro {a : Prop} (H : ¬a) : a ↔ false :=
iff.intro iff.intro
(assume H1 : a, absurd H1 H) (assume H1 : a, absurd H1 H)
(assume H2 : false, false_elim H2) (assume H2 : false, false.elim H2)
theorem a_neq_a {A : Type} (a : A) : (a ≠ a) ↔ false := theorem a_neq_a {A : Type} (a : A) : (a ≠ a) ↔ false :=
iff.intro iff.intro
(assume H, false.of_ne H) (assume H, false.of_ne H)
(assume H, false_elim H) (assume H, false.elim H)
theorem eq_id {A : Type} (a : A) : (a = a) ↔ true := theorem eq_id {A : Type} (a : A) : (a = a) ↔ true :=
iff_true_intro rfl iff_true_intro rfl
@ -134,7 +135,7 @@ iff.intro
(assume H, (assume H,
have H' : ¬a, from assume Ha, (H ▸ Ha) Ha, have H' : ¬a, from assume Ha, (H ▸ Ha) Ha,
H' (H⁻¹ ▸ H')) H' (H⁻¹ ▸ H'))
(assume H, false_elim H) (assume H, false.elim H)
theorem true_eq_false : (true ↔ false) ↔ false := theorem true_eq_false : (true ↔ false) ↔ false :=
not_true_iff_false ▸ (a_iff_not_a true) not_true_iff_false ▸ (a_iff_not_a true)

View file

@ -35,16 +35,16 @@ is_congruence2.mk
(take a1 b1 a2 b2, (take a1 b1 a2 b2,
assume H1 : a1 ↔ b1, assume H2 : a2 ↔ b2, assume H1 : a1 ↔ b1, assume H2 : a2 ↔ b2,
iff.intro iff.intro
(assume H3 : a1 ∧ a2, and.imp_and H3 (iff.elim_left H1) (iff.elim_left H2)) (assume H3 : a1 ∧ a2, and_of_and_of_imp_of_imp H3 (iff.elim_left H1) (iff.elim_left H2))
(assume H3 : b1 ∧ b2, and.imp_and H3 (iff.elim_right H1) (iff.elim_right H2))) (assume H3 : b1 ∧ b2, and_of_and_of_imp_of_imp H3 (iff.elim_right H1) (iff.elim_right H2)))
theorem is_congruence_or : is_congruence2 iff iff iff or := theorem is_congruence_or : is_congruence2 iff iff iff or :=
is_congruence2.mk is_congruence2.mk
(take a1 b1 a2 b2, (take a1 b1 a2 b2,
assume H1 : a1 ↔ b1, assume H2 : a2 ↔ b2, assume H1 : a1 ↔ b1, assume H2 : a2 ↔ b2,
iff.intro iff.intro
(assume H3 : a1 a2, or.imp_or H3 (iff.elim_left H1) (iff.elim_left H2)) (assume H3 : a1 a2, or_of_or_of_imp_of_imp H3 (iff.elim_left H1) (iff.elim_left H2))
(assume H3 : b1 b2, or.imp_or H3 (iff.elim_right H1) (iff.elim_right H2))) (assume H3 : b1 b2, or_of_or_of_imp_of_imp H3 (iff.elim_right H1) (iff.elim_right H2)))
theorem is_congruence_imp : is_congruence2 iff iff iff imp := theorem is_congruence_imp : is_congruence2 iff iff iff imp :=
is_congruence2.mk is_congruence2.mk
@ -87,8 +87,8 @@ relation.mp_like.mk (λa b (H : a ↔ b), iff.elim_left H)
/- support for calculations with iff -/ /- support for calculations with iff -/
namespace iff namespace iff
theorem subst {P : Prop → Prop} [C : is_congruence iff iff P] {a b : Prop} (H : a ↔ b) (H1 : P a) : theorem subst {P : Prop → Prop} [C : is_congruence iff iff P] {a b : Prop}
P b := (H : a ↔ b) (H1 : P a) : P b :=
@general_subst.subst Prop iff P C a b H H1 @general_subst.subst Prop iff P C a b H H1
end iff end iff

View file

@ -1,6 +1,11 @@
-- Copyright (c) 2014 Microsoft Corporation. All rights reserved. /-
-- Released under Apache 2.0 license as described in the file LICENSE. Copyright (c) 2014 Microsoft Corporation. All rights reserved.
-- Authors: Leonardo de Moura, Jeremy Avigad Released under Apache 2.0 license as described in the file LICENSE.
Module: logic.quantifiers
Authors: Leonardo de Moura, Jeremy Avigad
-/
open inhabited nonempty open inhabited nonempty
theorem exists_not_forall {A : Type} {p : A → Prop} (H : ∃x, p x) : ¬∀x, ¬p x := theorem exists_not_forall {A : Type} {p : A → Prop} (H : ∃x, p x) : ¬∀x, ¬p x :=

View file

@ -4,7 +4,7 @@ definition Prop := Type.{0}
definition false := ∀x : Prop, x definition false := ∀x : Prop, x
check false check false
theorem false_elim (C : Prop) (H : false) : C theorem false.elim (C : Prop) (H : false) : C
:= H C := H C
definition eq {A : Type} (a b : A) definition eq {A : Type} (a b : A)

View file

@ -4,7 +4,7 @@ definition Prop := Type.{0}
definition false : Prop := ∀x : Prop, x definition false : Prop := ∀x : Prop, x
check false check false
theorem false_elim (C : Prop) (H : false) : C theorem false.elim (C : Prop) (H : false) : C
:= H C := H C
definition eq {A : Type} (a b : A) definition eq {A : Type} (a b : A)

View file

@ -4,7 +4,7 @@ definition Prop := Type.{0}
definition false : Prop := ∀x : Prop, x definition false : Prop := ∀x : Prop, x
check false check false
theorem false_elim (C : Prop) (H : false) : C theorem false.elim (C : Prop) (H : false) : C
:= H C := H C
definition eq {A : Type} (a b : A) definition eq {A : Type} (a b : A)

View file

@ -14,7 +14,7 @@ theorem is_nil_nil (A : Type) : is_nil (@nil A)
:= of_eq_true (refl true) := of_eq_true (refl true)
theorem cons_ne_nil {A : Type} (a : A) (l : list A) : ¬ cons a l = nil theorem cons_ne_nil {A : Type} (a : A) (l : list A) : ¬ cons a l = nil
:= not_intro (assume H : cons a l = nil, := not.intro (assume H : cons a l = nil,
absurd absurd
(calc true = is_nil (@nil A) : refl _ (calc true = is_nil (@nil A) : refl _
... = is_nil (cons a l) : { symm H } ... = is_nil (cons a l) : { symm H }

View file

@ -14,7 +14,7 @@ section
notation `∅` := empty notation `∅` := empty
theorem mem_empty (x : T) : ¬ (x ∈ ∅) theorem mem_empty (x : T) : ¬ (x ∈ ∅)
:= not_intro (λH : x ∈ ∅, absurd H ff_ne_tt) := not.intro (λH : x ∈ ∅, absurd H ff_ne_tt)
end end
end set end set

View file

@ -48,13 +48,13 @@ rec_on s1
(take a2, show decidable (inl B a1 = inl B a2), from H1 a1 a2) (take a2, show decidable (inl B a1 = inl B a2), from H1 a1 a2)
(take b2, (take b2,
have H3 : (inl B a1 = inr A b2) ↔ false, have H3 : (inl B a1 = inr A b2) ↔ false,
from iff.intro inl_neq_inr (assume H4, false_elim H4), from iff.intro inl_neq_inr (assume H4, false.elim H4),
show decidable (inl B a1 = inr A b2), from decidable_iff_equiv _ (iff.symm H3))) show decidable (inl B a1 = inr A b2), from decidable_iff_equiv _ (iff.symm H3)))
(take b1, show decidable (inr A b1 = s2), from (take b1, show decidable (inr A b1 = s2), from
rec_on s2 rec_on s2
(take a2, (take a2,
have H3 : (inr A b1 = inl B a2) ↔ false, have H3 : (inr A b1 = inl B a2) ↔ false,
from iff.intro (assume H4, inl_neq_inr (symm H4)) (assume H4, false_elim H4), from iff.intro (assume H4, inl_neq_inr (symm H4)) (assume H4, false.elim H4),
show decidable (inr A b1 = inl B a2), from decidable_iff_equiv _ (iff.symm H3)) show decidable (inr A b1 = inl B a2), from decidable_iff_equiv _ (iff.symm H3))
(take b2, show decidable (inr A b1 = inr A b2), from H2 b1 b2)) (take b2, show decidable (inr A b1 = inr A b2), from H2 b1 b2))

View file

@ -60,7 +60,7 @@ theorem zero_or_succ (n : ) : n = 0 n = succ (pred n)
(show succ m = succ (pred (succ m)), from congr_arg succ (pred_succ m⁻¹))) (show succ m = succ (pred (succ m)), from congr_arg succ (pred_succ m⁻¹)))
theorem zero_or_succ2 (n : ) : n = 0 ∃k, n = succ k theorem zero_or_succ2 (n : ) : n = 0 ∃k, n = succ k
:= or.imp_or (zero_or_succ n) (assume H, H) (assume H : n = succ (pred n), exists_intro (pred n) H) := or_of_or_of_imp_of_imp (zero_or_succ n) (assume H, H) (assume H : n = succ (pred n), exists_intro (pred n) H)
theorem case {P : → Prop} (n : ) (H1: P 0) (H2 : ∀m, P (succ m)) : P n theorem case {P : → Prop} (n : ) (H1: P 0) (H2 : ∀m, P (succ m)) : P n
:= induction_on n H1 (take m IH, H2 m) := induction_on n H1 (take m IH, H2 m)
@ -502,10 +502,10 @@ theorem succ_le_left_or {n m : } (H : n ≤ m) : succ n ≤ m n = m
or_intro_left _ Hlt) or_intro_left _ Hlt)
theorem succ_le_left {n m : } (H1 : n ≤ m) (H2 : n ≠ m) : succ n ≤ m theorem succ_le_left {n m : } (H1 : n ≤ m) (H2 : n ≠ m) : succ n ≤ m
:= or.resolve_left (succ_le_left_or H1) H2 := or_resolve_left (succ_le_left_or H1) H2
theorem succ_le_right_inv {n m : } (H : n ≤ succ m) : n ≤ m n = succ m theorem succ_le_right_inv {n m : } (H : n ≤ succ m) : n ≤ m n = succ m
:= or.imp_or (succ_le_left_or H) := or_of_or_of_imp_of_imp (succ_le_left_or H)
(take H2 : succ n ≤ succ m, show n ≤ m, from succ_le_cancel H2) (take H2 : succ n ≤ succ m, show n ≤ m, from succ_le_cancel H2)
(take H2 : n = succ m, H2) (take H2 : n = succ m, H2)
@ -561,7 +561,7 @@ theorem pred_le_left_inv {n m : } (H : pred n ≤ m) : n ≤ m n = succ m
have H3 : k ≤ m, from subst H2 H, have H3 : k ≤ m, from subst H2 H,
have H4 : succ k ≤ m k = m, from succ_le_left_or H3, have H4 : succ k ≤ m k = m, from succ_le_left_or H3,
show n ≤ m n = succ m, from show n ≤ m n = succ m, from
or.imp_or H4 or_of_or_of_imp_of_imp H4
(take H5 : succ k ≤ m, show n ≤ m, from subst (symm Hn) H5) (take H5 : succ k ≤ m, show n ≤ m, from subst (symm Hn) H5)
(take H5 : k = m, show n = succ m, from subst H5 Hn)) (take H5 : k = m, show n = succ m, from subst H5 Hn))
@ -589,7 +589,7 @@ theorem le_imp_succ_le_or_eq {n m : } (H : n ≤ m) : succ n ≤ m n = m
or_intro_left _ Hlt) or_intro_left _ Hlt)
theorem le_ne_imp_succ_le {n m : } (H1 : n ≤ m) (H2 : n ≠ m) : succ n ≤ m theorem le_ne_imp_succ_le {n m : } (H1 : n ≤ m) (H2 : n ≠ m) : succ n ≤ m
:= or.resolve_left (le_imp_succ_le_or_eq H1) H2 := or_resolve_left (le_imp_succ_le_or_eq H1) H2
theorem succ_le_imp_le_and_ne {n m : } (H : succ n ≤ m) : n ≤ m ∧ n ≠ m theorem succ_le_imp_le_and_ne {n m : } (H : succ n ≤ m) : n ≤ m ∧ n ≠ m
:= :=
@ -620,7 +620,7 @@ theorem pred_le_imp_le_or_eq {n m : } (H : pred n ≤ m) : n ≤ m n = su
have H3 : k ≤ m, from subst H2 H, have H3 : k ≤ m, from subst H2 H,
have H4 : succ k ≤ m k = m, from le_imp_succ_le_or_eq H3, have H4 : succ k ≤ m k = m, from le_imp_succ_le_or_eq H3,
show n ≤ m n = succ m, from show n ≤ m n = succ m, from
or.imp_or H4 or_of_or_of_imp_of_imp H4
(take H5 : succ k ≤ m, show n ≤ m, from subst (symm Hn) H5) (take H5 : succ k ≤ m, show n ≤ m, from subst (symm Hn) H5)
(take H5 : k = m, show n = succ m, from subst H5 Hn)) (take H5 : k = m, show n = succ m, from subst H5 Hn))
@ -806,13 +806,13 @@ theorem le_or_lt (n m : ) : n ≤ m m < n
(assume H : m < k, or_intro_right _ (succ_lt_right H))) (assume H : m < k, or_intro_right _ (succ_lt_right H)))
theorem trichotomy_alt (n m : ) : (n < m n = m) m < n theorem trichotomy_alt (n m : ) : (n < m n = m) m < n
:= or.imp_or (le_or_lt n m) (assume H : n ≤ m, le_imp_lt_or_eq H) (assume H : m < n, H) := or_of_or_of_imp_of_imp (le_or_lt n m) (assume H : n ≤ m, le_imp_lt_or_eq H) (assume H : m < n, H)
theorem trichotomy (n m : ) : n < m n = m m < n theorem trichotomy (n m : ) : n < m n = m m < n
:= iff.elim_left or.assoc (trichotomy_alt n m) := iff.elim_left or.assoc (trichotomy_alt n m)
theorem le_total (n m : ) : n ≤ m m ≤ n theorem le_total (n m : ) : n ≤ m m ≤ n
:= or.imp_or (le_or_lt n m) (assume H : n ≤ m, H) (assume H : m < n, lt_imp_le H) := or_of_or_of_imp_of_imp (le_or_lt n m) (assume H : n ≤ m, H) (assume H : m < n, lt_imp_le H)
-- interaction with mul under "positivity" -- interaction with mul under "positivity"
@ -871,7 +871,7 @@ theorem add_eq_self {n m : } (H : n + m = n) : m = 0
---------- basic ---------- basic
theorem zero_or_positive (n : ) : n = 0 n > 0 theorem zero_or_positive (n : ) : n = 0 n > 0
:= or.imp_or (or.swap (le_imp_lt_or_eq (zero_le n))) (take H : 0 = n, symm H) (take H : n > 0, H) := or_of_or_of_imp_of_imp (or.swap (le_imp_lt_or_eq (zero_le n))) (take H : 0 = n, symm H) (take H : n > 0, H)
theorem succ_positive {n m : } (H : n = succ m) : n > 0 theorem succ_positive {n m : } (H : n = succ m) : n > 0
:= subst (symm H) (lt_zero m) := subst (symm H) (lt_zero m)
@ -954,7 +954,7 @@ theorem mul_left_inj {n m k : } (Hn : n > 0) (H : n * m = n * k) : m = k
n * m = n * 0 : H n * m = n * 0 : H
... = 0 : mul_zero_right n, ... = 0 : mul_zero_right n,
have H3 : n = 0 m = 0, from mul_eq_zero H2, have H3 : n = 0 m = 0, from mul_eq_zero H2,
or.resolve_right H3 (ne.symm (lt_ne Hn))) or_resolve_right H3 (ne.symm (lt_ne Hn)))
(take (l : ), (take (l : ),
assume (IH : ∀ m, n * m = n * l → m = l), assume (IH : ∀ m, n * m = n * l → m = l),
take (m : ), take (m : ),

View file

@ -54,7 +54,7 @@ theorem zero_or_succ (n : ) : n = 0 n = succ (pred n)
(show succ m = succ (pred (succ m)), from congr_arg succ (pred_succ m⁻¹))) (show succ m = succ (pred (succ m)), from congr_arg succ (pred_succ m⁻¹)))
theorem zero_or_succ2 (n : ) : n = 0 ∃k, n = succ k theorem zero_or_succ2 (n : ) : n = 0 ∃k, n = succ k
:= or.imp_or (zero_or_succ n) (assume H, H) (assume H : n = succ (pred n), exists_intro (pred n) H) := or_of_or_of_imp_of_imp (zero_or_succ n) (assume H, H) (assume H : n = succ (pred n), exists_intro (pred n) H)
theorem case {P : → Prop} (n : ) (H1: P 0) (H2 : ∀m, P (succ m)) : P n theorem case {P : → Prop} (n : ) (H1: P 0) (H2 : ∀m, P (succ m)) : P n
:= induction_on n H1 (take m IH, H2 m) := induction_on n H1 (take m IH, H2 m)
@ -496,10 +496,10 @@ theorem succ_le_left_or {n m : } (H : n ≤ m) : succ n ≤ m n = m
or.intro_left _ Hlt) or.intro_left _ Hlt)
theorem succ_le_left {n m : } (H1 : n ≤ m) (H2 : n ≠ m) : succ n ≤ m theorem succ_le_left {n m : } (H1 : n ≤ m) (H2 : n ≠ m) : succ n ≤ m
:= or.resolve_left (succ_le_left_or H1) H2 := or_resolve_left (succ_le_left_or H1) H2
theorem succ_le_right_inv {n m : } (H : n ≤ succ m) : n ≤ m n = succ m theorem succ_le_right_inv {n m : } (H : n ≤ succ m) : n ≤ m n = succ m
:= or.imp_or (succ_le_left_or H) := or_of_or_of_imp_of_imp (succ_le_left_or H)
(take H2 : succ n ≤ succ m, show n ≤ m, from succ_le_cancel H2) (take H2 : succ n ≤ succ m, show n ≤ m, from succ_le_cancel H2)
(take H2 : n = succ m, H2) (take H2 : n = succ m, H2)
@ -555,7 +555,7 @@ theorem pred_le_left_inv {n m : } (H : pred n ≤ m) : n ≤ m n = succ m
have H3 : k ≤ m, from subst H2 H, have H3 : k ≤ m, from subst H2 H,
have H4 : succ k ≤ m k = m, from succ_le_left_or H3, have H4 : succ k ≤ m k = m, from succ_le_left_or H3,
show n ≤ m n = succ m, from show n ≤ m n = succ m, from
or.imp_or H4 or_of_or_of_imp_of_imp H4
(take H5 : succ k ≤ m, show n ≤ m, from subst (symm Hn) H5) (take H5 : succ k ≤ m, show n ≤ m, from subst (symm Hn) H5)
(take H5 : k = m, show n = succ m, from subst H5 Hn)) (take H5 : k = m, show n = succ m, from subst H5 Hn))
@ -583,10 +583,10 @@ theorem le_imp_succ_le_or_eq {n m : } (H : n ≤ m) : succ n ≤ m n = m
or.intro_left _ Hlt) or.intro_left _ Hlt)
theorem le_ne_imp_succ_le {n m : } (H1 : n ≤ m) (H2 : n ≠ m) : succ n ≤ m theorem le_ne_imp_succ_le {n m : } (H1 : n ≤ m) (H2 : n ≠ m) : succ n ≤ m
:= or.resolve_left (le_imp_succ_le_or_eq H1) H2 := or_resolve_left (le_imp_succ_le_or_eq H1) H2
theorem le_succ_imp_le_or_eq {n m : } (H : n ≤ succ m) : n ≤ m n = succ m theorem le_succ_imp_le_or_eq {n m : } (H : n ≤ succ m) : n ≤ m n = succ m
:= or.imp_or_left (le_imp_succ_le_or_eq H) := or_of_or_of_imp_left (le_imp_succ_le_or_eq H)
(take H2 : succ n ≤ succ m, show n ≤ m, from succ_le_cancel H2) (take H2 : succ n ≤ succ m, show n ≤ m, from succ_le_cancel H2)
theorem succ_le_imp_le_and_ne {n m : } (H : succ n ≤ m) : n ≤ m ∧ n ≠ m theorem succ_le_imp_le_and_ne {n m : } (H : succ n ≤ m) : n ≤ m ∧ n ≠ m
@ -618,7 +618,7 @@ theorem pred_le_imp_le_or_eq {n m : } (H : pred n ≤ m) : n ≤ m n = su
have H3 : k ≤ m, from subst H2 H, have H3 : k ≤ m, from subst H2 H,
have H4 : succ k ≤ m k = m, from le_imp_succ_le_or_eq H3, have H4 : succ k ≤ m k = m, from le_imp_succ_le_or_eq H3,
show n ≤ m n = succ m, from show n ≤ m n = succ m, from
or.imp_or H4 or_of_or_of_imp_of_imp H4
(take H5 : succ k ≤ m, show n ≤ m, from subst (symm Hn) H5) (take H5 : succ k ≤ m, show n ≤ m, from subst (symm Hn) H5)
(take H5 : k = m, show n = succ m, from subst H5 Hn)) (take H5 : k = m, show n = succ m, from subst H5 Hn))
@ -804,13 +804,13 @@ theorem le_or_lt (n m : ) : n ≤ m m < n
(assume H : m < k, or.intro_right _ (succ_lt_right H))) (assume H : m < k, or.intro_right _ (succ_lt_right H)))
theorem trichotomy_alt (n m : ) : (n < m n = m) m < n theorem trichotomy_alt (n m : ) : (n < m n = m) m < n
:= or.imp_or (le_or_lt n m) (assume H : n ≤ m, le_imp_lt_or_eq H) (assume H : m < n, H) := or_of_or_of_imp_of_imp (le_or_lt n m) (assume H : n ≤ m, le_imp_lt_or_eq H) (assume H : m < n, H)
theorem trichotomy (n m : ) : n < m n = m m < n theorem trichotomy (n m : ) : n < m n = m m < n
:= iff.elim_left or.assoc (trichotomy_alt n m) := iff.elim_left or.assoc (trichotomy_alt n m)
theorem le_total (n m : ) : n ≤ m m ≤ n theorem le_total (n m : ) : n ≤ m m ≤ n
:= or.imp_or (le_or_lt n m) (assume H : n ≤ m, H) (assume H : m < n, lt_imp_le H) := or_of_or_of_imp_of_imp (le_or_lt n m) (assume H : n ≤ m, H) (assume H : m < n, lt_imp_le H)
-- interaction with mul under "positivity" -- interaction with mul under "positivity"
@ -869,7 +869,7 @@ theorem add_eq_self {n m : } (H : n + m = n) : m = 0
---------- basic ---------- basic
theorem zero_or_positive (n : ) : n = 0 n > 0 theorem zero_or_positive (n : ) : n = 0 n > 0
:= or.imp_or (or.swap (le_imp_lt_or_eq (zero_le n))) (take H : 0 = n, symm H) (take H : n > 0, H) := or_of_or_of_imp_of_imp (or.swap (le_imp_lt_or_eq (zero_le n))) (take H : 0 = n, symm H) (take H : n > 0, H)
theorem succ_positive {n m : } (H : n = succ m) : n > 0 theorem succ_positive {n m : } (H : n = succ m) : n > 0
:= subst (symm H) (lt_zero m) := subst (symm H) (lt_zero m)
@ -906,7 +906,7 @@ theorem case_zero_pos {P : → Prop} (y : ) (H0 : P 0) (H1 : ∀y, y > 0
:= case y H0 (take y', H1 _ (succ_pos _)) := case y H0 (take y', H1 _ (succ_pos _))
theorem zero_or_pos (n : ) : n = 0 n > 0 theorem zero_or_pos (n : ) : n = 0 n > 0
:= or.imp_or_left (or.swap (le_imp_lt_or_eq (zero_le n))) (take H : 0 = n, symm H) := or_of_or_of_imp_left (or.swap (le_imp_lt_or_eq (zero_le n))) (take H : 0 = n, symm H)
theorem succ_imp_pos {n m : } (H : n = succ m) : n > 0 theorem succ_imp_pos {n m : } (H : n = succ m) : n > 0
:= subst (symm H) (succ_pos m) := subst (symm H) (succ_pos m)
@ -958,7 +958,7 @@ theorem mul_left_inj {n m k : } (Hn : n > 0) (H : n * m = n * k) : m = k
n * m = n * 0 : H n * m = n * 0 : H
... = 0 : mul_zero_right n, ... = 0 : mul_zero_right n,
have H3 : n = 0 m = 0, from mul_eq_zero H2, have H3 : n = 0 m = 0, from mul_eq_zero H2,
or.resolve_right H3 (ne.symm (lt_ne Hn))) or_resolve_right H3 (ne.symm (lt_ne Hn)))
(take (l : ), (take (l : ),
assume (IH : ∀ m, n * m = n * l → m = l), assume (IH : ∀ m, n * m = n * l → m = l),
take (m : ), take (m : ),