feat(frontends/lean): round parenthesis for [tactic1 | tactic2]

This commit also replaces the notation for divides
     `(` a `|` b `)`
with
      a `∣` b

The character `∣` is entered by typing \|

closes #516
This commit is contained in:
Leonardo de Moura 2015-04-06 09:24:09 -07:00
parent 969d17fd12
commit 754276a660
25 changed files with 168 additions and 175 deletions

View file

@ -93,7 +93,7 @@ namespace functor
fapply @is_iso.mk, apply (F (f⁻¹)), fapply @is_iso.mk, apply (F (f⁻¹)),
repeat (apply concat ; apply inverse ; apply (respect_comp F) ; repeat (apply concat ; apply inverse ; apply (respect_comp F) ;
apply concat ; apply (ap (λ x, to_fun_hom F x)) ; apply concat ; apply (ap (λ x, to_fun_hom F x)) ;
[apply left_inverse | apply right_inverse] ; (apply left_inverse | apply right_inverse);
apply (respect_id F) ), apply (respect_id F) ),
end end

View file

@ -92,7 +92,7 @@ reserve infixl ``:65
/- other symbols -/ /- other symbols -/
reserve notation `(` a `|` b `)` reserve infix ``:50
reserve infixl `++`:65 reserve infixl `++`:65
reserve infixr `::`:65 reserve infixr `::`:65

View file

@ -92,13 +92,13 @@ opaque definition change (e : expr) : tactic := builtin
opaque definition assert_hypothesis (id : expr) (e : expr) : tactic := builtin opaque definition assert_hypothesis (id : expr) (e : expr) : tactic := builtin
infixl `;`:15 := and_then infixl `;`:15 := and_then
notation `[` h:10 `|`:10 r:(foldl:10 `|` (e r, or_else r e) h) `]` := r notation `(` h `|` r:(foldl `|` (e r, or_else r e) h) `)` := r
definition try (t : tactic) : tactic := [t | id] definition try (t : tactic) : tactic := (t | id)
definition repeat1 (t : tactic) : tactic := t ; repeat t definition repeat1 (t : tactic) : tactic := t ; repeat t
definition focus (t : tactic) : tactic := focus_at t 0 definition focus (t : tactic) : tactic := focus_at t 0
definition determ (t : tactic) : tactic := at_most t 1 definition determ (t : tactic) : tactic := at_most t 1
definition trivial : tactic := [ apply eq.refl | assumption ] definition trivial : tactic := ( apply eq.refl | assumption )
definition do (n : num) (t : tactic) : tactic := definition do (n : num) (t : tactic) : tactic :=
nat.rec id (λn t', (t;t')) (nat.of_num n) nat.rec id (λn t', (t;t')) (nat.of_num n)

View file

@ -544,10 +544,10 @@ section
... = -1 * abs a : by rewrite neg_eq_neg_one_mul ... = -1 * abs a : by rewrite neg_eq_neg_one_mul
... = sign a * abs a : by rewrite (sign_of_neg H1)) ... = sign a * abs a : by rewrite (sign_of_neg H1))
theorem abs_dvd_iff_dvd (a b : A) : (abs a | b) ↔ (a | b) := theorem abs_dvd_iff_dvd (a b : A) : abs a b ↔ a b :=
abs.by_cases !iff.refl !neg_dvd_iff_dvd abs.by_cases !iff.refl !neg_dvd_iff_dvd
theorem dvd_abs_iff (a b : A) : (a | abs b) ↔ (a | b) := theorem dvd_abs_iff (a b : A) : a abs b ↔ a b :=
abs.by_cases !iff.refl !dvd_neg_iff_dvd abs.by_cases !iff.refl !dvd_neg_iff_dvd
theorem abs_mul (a b : A) : abs (a * b) = abs a * abs b := theorem abs_mul (a b : A) : abs (a * b) = abs a * abs b :=

View file

@ -72,28 +72,28 @@ section comm_semiring
include s include s
definition dvd (a b : A) : Prop := ∃c, b = a * c definition dvd (a b : A) : Prop := ∃c, b = a * c
notation ( a | b ) := dvd a b notation a b := dvd a b
theorem dvd.intro {a b c : A} (H : a * c = b) : (a | b) := theorem dvd.intro {a b c : A} (H : a * c = b) : a b :=
exists.intro _ H⁻¹ exists.intro _ H⁻¹
theorem dvd.intro_left {a b c : A} (H : c * a = b) : (a | b) := theorem dvd.intro_left {a b c : A} (H : c * a = b) : a b :=
dvd.intro (!mul.comm ▸ H) dvd.intro (!mul.comm ▸ H)
theorem exists_eq_mul_right_of_dvd {a b : A} (H : (a | b)) : ∃c, b = a * c := H theorem exists_eq_mul_right_of_dvd {a b : A} (H : a b) : ∃c, b = a * c := H
theorem dvd.elim {P : Prop} {a b : A} (H₁ : (a | b)) (H₂ : ∀c, b = a * c → P) : P := theorem dvd.elim {P : Prop} {a b : A} (H₁ : a b) (H₂ : ∀c, b = a * c → P) : P :=
exists.elim H₁ H₂ exists.elim H₁ H₂
theorem exists_eq_mul_left_of_dvd {a b : A} (H : (a | b)) : ∃c, b = c * a := theorem exists_eq_mul_left_of_dvd {a b : A} (H : a b) : ∃c, b = c * a :=
dvd.elim H (take c, assume H1 : b = a * c, exists.intro c (H1 ⬝ !mul.comm)) dvd.elim H (take c, assume H1 : b = a * c, exists.intro c (H1 ⬝ !mul.comm))
theorem dvd.elim_left {P : Prop} {a b : A} (H₁ : (a | b)) (H₂ : ∀c, b = c * a → P) : P := theorem dvd.elim_left {P : Prop} {a b : A} (H₁ : a b) (H₂ : ∀c, b = c * a → P) : P :=
exists.elim (exists_eq_mul_left_of_dvd H₁) (take c, assume H₃ : b = c * a, H₂ c H₃) exists.elim (exists_eq_mul_left_of_dvd H₁) (take c, assume H₃ : b = c * a, H₂ c H₃)
theorem dvd.refl : (a | a) := dvd.intro !mul_one theorem dvd.refl : a a := dvd.intro !mul_one
theorem dvd.trans {a b c : A} (H₁ : (a | b)) (H₂ : (b | c)) : (a | c) := theorem dvd.trans {a b c : A} (H₁ : a b) (H₂ : b c) : a c :=
dvd.elim H₁ dvd.elim H₁
(take d, assume H₃ : b = a * d, (take d, assume H₃ : b = a * d,
dvd.elim H₂ dvd.elim H₂
@ -101,28 +101,28 @@ section comm_semiring
dvd.intro dvd.intro
(show a * (d * e) = c, by rewrite [-mul.assoc, -H₃, H₄]))) (show a * (d * e) = c, by rewrite [-mul.assoc, -H₃, H₄])))
theorem eq_zero_of_zero_dvd {a : A} (H : (0 | a)) : a = 0 := theorem eq_zero_of_zero_dvd {a : A} (H : 0 a) : a = 0 :=
dvd.elim H (take c, assume H' : a = 0 * c, H' ⬝ !zero_mul) dvd.elim H (take c, assume H' : a = 0 * c, H' ⬝ !zero_mul)
theorem dvd_zero : (a | 0) := dvd.intro !mul_zero theorem dvd_zero : a 0 := dvd.intro !mul_zero
theorem one_dvd : (1 | a) := dvd.intro !one_mul theorem one_dvd : 1 a := dvd.intro !one_mul
theorem dvd_mul_right : (a | a * b) := dvd.intro rfl theorem dvd_mul_right : a a * b := dvd.intro rfl
theorem dvd_mul_left : (a | b * a) := mul.comm a b ▸ dvd_mul_right a b theorem dvd_mul_left : a b * a := mul.comm a b ▸ dvd_mul_right a b
theorem dvd_mul_of_dvd_left {a b : A} (H : (a | b)) (c : A) : (a | b * c) := theorem dvd_mul_of_dvd_left {a b : A} (H : a b) (c : A) : a b * c :=
dvd.elim H dvd.elim H
(take d, (take d,
assume H₁ : b = a * d, assume H₁ : b = a * d,
dvd.intro dvd.intro
(show a * (d * c) = b * c, from by rewrite [-mul.assoc, H₁])) (show a * (d * c) = b * c, from by rewrite [-mul.assoc, H₁]))
theorem dvd_mul_of_dvd_right {a b : A} (H : (a | b)) (c : A) : (a | c * b) := theorem dvd_mul_of_dvd_right {a b : A} (H : a b) (c : A) : a c * b :=
!mul.comm ▸ (dvd_mul_of_dvd_left H _) !mul.comm ▸ (dvd_mul_of_dvd_left H _)
theorem mul_dvd_mul {a b c d : A} (dvd_ab : (a | b)) (dvd_cd : (c | d)) : (a * c | b * d) := theorem mul_dvd_mul {a b c d : A} (dvd_ab : (a b)) (dvd_cd : c d) : a * c b * d :=
dvd.elim dvd_ab dvd.elim dvd_ab
(take e, assume Haeb : b = a * e, (take e, assume Haeb : b = a * e,
dvd.elim dvd_cd dvd.elim dvd_cd
@ -131,13 +131,13 @@ section comm_semiring
(show a * c * (e * f) = b * d, (show a * c * (e * f) = b * d,
by rewrite [mul.assoc, {c*_}mul.left_comm, -mul.assoc, Haeb, Hcfd]))) by rewrite [mul.assoc, {c*_}mul.left_comm, -mul.assoc, Haeb, Hcfd])))
theorem dvd_of_mul_right_dvd {a b c : A} (H : (a * b | c)) : (a | c) := theorem dvd_of_mul_right_dvd {a b c : A} (H : a * b c) : a c :=
dvd.elim H (take d, assume Habdc : c = a * b * d, dvd.intro (!mul.assoc⁻¹ ⬝ Habdc⁻¹)) dvd.elim H (take d, assume Habdc : c = a * b * d, dvd.intro (!mul.assoc⁻¹ ⬝ Habdc⁻¹))
theorem dvd_of_mul_left_dvd {a b c : A} (H : (a * b | c)) : (b | c) := theorem dvd_of_mul_left_dvd {a b c : A} (H : a * b c) : b c :=
dvd_of_mul_right_dvd (mul.comm a b ▸ H) dvd_of_mul_right_dvd (mul.comm a b ▸ H)
theorem dvd_add {a b c : A} (Hab : (a | b)) (Hac : (a | c)) : (a | b + c) := theorem dvd_add {a b c : A} (Hab : a b) (Hac : a c) : a b + c :=
dvd.elim Hab dvd.elim Hab
(take d, assume Hadb : b = a * d, (take d, assume Hadb : b = a * d,
dvd.elim Hac dvd.elim Hac
@ -259,35 +259,35 @@ section
theorem mul_self_sub_one_eq : a * a - 1 = (a + 1) * (a - 1) := theorem mul_self_sub_one_eq : a * a - 1 = (a + 1) * (a - 1) :=
mul_one 1 ▸ mul_self_sub_mul_self_eq a 1 mul_one 1 ▸ mul_self_sub_mul_self_eq a 1
theorem dvd_neg_iff_dvd : (a | -b) ↔ (a | b) := theorem dvd_neg_iff_dvd : (a -b) ↔ (a b) :=
iff.intro iff.intro
(assume H : (a | -b), (assume H : (a -b),
dvd.elim H dvd.elim H
(take c, assume H' : -b = a * c, (take c, assume H' : -b = a * c,
dvd.intro dvd.intro
(show a * -c = b, (show a * -c = b,
by rewrite [-neg_mul_eq_mul_neg, -H', neg_neg]))) by rewrite [-neg_mul_eq_mul_neg, -H', neg_neg])))
(assume H : (a | b), (assume H : (a b),
dvd.elim H dvd.elim H
(take c, assume H' : b = a * c, (take c, assume H' : b = a * c,
dvd.intro dvd.intro
(show a * -c = -b, (show a * -c = -b,
by rewrite [-neg_mul_eq_mul_neg, -H']))) by rewrite [-neg_mul_eq_mul_neg, -H'])))
theorem neg_dvd_iff_dvd : (-a | b) ↔ (a | b) := theorem neg_dvd_iff_dvd : (-a b) ↔ (a b) :=
iff.intro iff.intro
(assume H : (-a | b), (assume H : (-a b),
dvd.elim H dvd.elim H
(take c, assume H' : b = -a * c, (take c, assume H' : b = -a * c,
dvd.intro dvd.intro
(show a * -c = b, by rewrite [-neg_mul_comm, H']))) (show a * -c = b, by rewrite [-neg_mul_comm, H'])))
(assume H : (a | b), (assume H : (a b),
dvd.elim H dvd.elim H
(take c, assume H' : b = a * c, (take c, assume H' : b = a * c,
dvd.intro dvd.intro
(show -a * -c = b, by rewrite [neg_mul_neg, H']))) (show -a * -c = b, by rewrite [neg_mul_neg, H'])))
theorem dvd_sub (H₁ : (a | b)) (H₂ : (a | c)) : (a | b - c) := theorem dvd_sub (H₁ : (a b)) (H₂ : (a c)) : (a b - c) :=
dvd_add H₁ (iff.elim_right !dvd_neg_iff_dvd H₂) dvd_add H₁ (iff.elim_right !dvd_neg_iff_dvd H₂)
end end
@ -343,14 +343,14 @@ section
-- TODO: c - b * c → c = 0 b = 1 and variants -- TODO: c - b * c → c = 0 b = 1 and variants
theorem dvd_of_mul_dvd_mul_left {a b c : A} (Ha : a ≠ 0) (Hdvd : (a * b | a * c)) : (b | c) := theorem dvd_of_mul_dvd_mul_left {a b c : A} (Ha : a ≠ 0) (Hdvd : (a * b a * c)) : (b c) :=
dvd.elim Hdvd dvd.elim Hdvd
(take d, (take d,
assume H : a * c = a * b * d, assume H : a * c = a * b * d,
have H1 : b * d = c, from mul.cancel_left Ha (mul.assoc a b d ▸ H⁻¹), have H1 : b * d = c, from mul.cancel_left Ha (mul.assoc a b d ▸ H⁻¹),
dvd.intro H1) dvd.intro H1)
theorem dvd_of_mul_dvd_mul_right {a b c : A} (Ha : a ≠ 0) (Hdvd : (b * a | c * a)) : (b | c) := theorem dvd_of_mul_dvd_mul_right {a b c : A} (Ha : a ≠ 0) (Hdvd : (b * a c * a)) : (b c) :=
dvd.elim Hdvd dvd.elim Hdvd
(take d, (take d,
assume H : c * a = b * a * d, assume H : c * a = b * a * d,

View file

@ -710,35 +710,35 @@ section port_algebra
theorem ne_zero_of_mul_ne_zero_left : ∀{a b : }, a * b ≠ 0 → b ≠ 0 := theorem ne_zero_of_mul_ne_zero_left : ∀{a b : }, a * b ≠ 0 → b ≠ 0 :=
@algebra.ne_zero_of_mul_ne_zero_left _ _ @algebra.ne_zero_of_mul_ne_zero_left _ _
definition dvd (a b : ) : Prop := algebra.dvd a b definition dvd (a b : ) : Prop := algebra.dvd a b
notation (a | b) := dvd a b notation a b := dvd a b
theorem dvd.intro : ∀{a b c : } (H : a * c = b), (a | b) := @algebra.dvd.intro _ _ theorem dvd.intro : ∀{a b c : } (H : a * c = b), a b := @algebra.dvd.intro _ _
theorem dvd.intro_left : ∀{a b c : } (H : c * a = b), (a | b) := @algebra.dvd.intro_left _ _ theorem dvd.intro_left : ∀{a b c : } (H : c * a = b), a b := @algebra.dvd.intro_left _ _
theorem exists_eq_mul_right_of_dvd : ∀{a b : } (H : (a | b)), ∃c, b = a * c := theorem exists_eq_mul_right_of_dvd : ∀{a b : } (H : a b), ∃c, b = a * c :=
@algebra.exists_eq_mul_right_of_dvd _ _ @algebra.exists_eq_mul_right_of_dvd _ _
theorem dvd.elim : ∀{P : Prop} {a b : } (H₁ : (a | b)) (H₂ : ∀c, b = a * c → P), P := theorem dvd.elim : ∀{P : Prop} {a b : } (H₁ : a b) (H₂ : ∀c, b = a * c → P), P :=
@algebra.dvd.elim _ _ @algebra.dvd.elim _ _
theorem exists_eq_mul_left_of_dvd : ∀{a b : } (H : (a | b)), ∃c, b = c * a := theorem exists_eq_mul_left_of_dvd : ∀{a b : } (H : a b), ∃c, b = c * a :=
@algebra.exists_eq_mul_left_of_dvd _ _ @algebra.exists_eq_mul_left_of_dvd _ _
theorem dvd.elim_left : ∀{P : Prop} {a b : } (H₁ : (a | b)) (H₂ : ∀c, b = c * a → P), P := theorem dvd.elim_left : ∀{P : Prop} {a b : } (H₁ : a b) (H₂ : ∀c, b = c * a → P), P :=
@algebra.dvd.elim_left _ _ @algebra.dvd.elim_left _ _
theorem dvd.refl : ∀a : , (a | a) := algebra.dvd.refl theorem dvd.refl : ∀a : , (a a) := algebra.dvd.refl
theorem dvd.trans : ∀{a b c : } (H₁ : (a | b)) (H₂ : (b | c)), (a | c) := @algebra.dvd.trans _ _ theorem dvd.trans : ∀{a b c : } (H₁ : a b) (H₂ : b c), a c := @algebra.dvd.trans _ _
theorem eq_zero_of_zero_dvd : ∀{a : } (H : (0 | a)), a = 0 := @algebra.eq_zero_of_zero_dvd _ _ theorem eq_zero_of_zero_dvd : ∀{a : } (H : 0 a), a = 0 := @algebra.eq_zero_of_zero_dvd _ _
theorem dvd_zero : ∀a : , (a | 0) := algebra.dvd_zero theorem dvd_zero : ∀a : , a 0 := algebra.dvd_zero
theorem one_dvd : ∀a : , (1 | a) := algebra.one_dvd theorem one_dvd : ∀a : , 1 a := algebra.one_dvd
theorem dvd_mul_right : ∀a b : , (a | a * b) := algebra.dvd_mul_right theorem dvd_mul_right : ∀a b : , a a * b := algebra.dvd_mul_right
theorem dvd_mul_left : ∀a b : , (a | b * a) := algebra.dvd_mul_left theorem dvd_mul_left : ∀a b : , a b * a := algebra.dvd_mul_left
theorem dvd_mul_of_dvd_left : ∀{a b : } (H : (a | b)) (c : ), (a | b * c) := theorem dvd_mul_of_dvd_left : ∀{a b : } (H : a b) (c : ), a b * c :=
@algebra.dvd_mul_of_dvd_left _ _ @algebra.dvd_mul_of_dvd_left _ _
theorem dvd_mul_of_dvd_right : ∀{a b : } (H : (a | b)) (c : ), (a | c * b) := theorem dvd_mul_of_dvd_right : ∀{a b : } (H : a b) (c : ), a c * b :=
@algebra.dvd_mul_of_dvd_right _ _ @algebra.dvd_mul_of_dvd_right _ _
theorem mul_dvd_mul : ∀{a b c d : }, (a | b) → (c | d) → (a * c | b * d) := theorem mul_dvd_mul : ∀{a b c d : }, a b → c d → a * c b * d :=
@algebra.mul_dvd_mul _ _ @algebra.mul_dvd_mul _ _
theorem dvd_of_mul_right_dvd : ∀{a b c : }, (a * b | c) → (a | c) := theorem dvd_of_mul_right_dvd : ∀{a b c : }, a * b c → a c :=
@algebra.dvd_of_mul_right_dvd _ _ @algebra.dvd_of_mul_right_dvd _ _
theorem dvd_of_mul_left_dvd : ∀{a b c : }, (a * b | c) → (b | c) := theorem dvd_of_mul_left_dvd : ∀{a b c : }, a * b c → b c :=
@algebra.dvd_of_mul_left_dvd _ _ @algebra.dvd_of_mul_left_dvd _ _
theorem dvd_add : ∀{a b c : }, (a | b) → (a | c) → (a | b + c) := @algebra.dvd_add _ _ theorem dvd_add : ∀{a b c : }, a b → a c → a b + c := @algebra.dvd_add _ _
theorem zero_mul : ∀a : , 0 * a = 0 := algebra.zero_mul theorem zero_mul : ∀a : , 0 * a = 0 := algebra.zero_mul
theorem mul_zero : ∀a : , a * 0 = 0 := algebra.mul_zero theorem mul_zero : ∀a : , a * 0 = 0 := algebra.mul_zero
theorem neg_mul_eq_neg_mul : ∀a b : , -(a * b) = -a * b := algebra.neg_mul_eq_neg_mul theorem neg_mul_eq_neg_mul : ∀a b : , -(a * b) = -a * b := algebra.neg_mul_eq_neg_mul
@ -757,9 +757,9 @@ section port_algebra
algebra.mul_self_sub_mul_self_eq algebra.mul_self_sub_mul_self_eq
theorem mul_self_sub_one_eq : ∀a : , a * a - 1 = (a + 1) * (a - 1) := theorem mul_self_sub_one_eq : ∀a : , a * a - 1 = (a + 1) * (a - 1) :=
algebra.mul_self_sub_one_eq algebra.mul_self_sub_one_eq
theorem dvd_neg_iff_dvd : ∀a b : , (a | -b) ↔ (a | b) := algebra.dvd_neg_iff_dvd theorem dvd_neg_iff_dvd : ∀a b : , a -b ↔ a b := algebra.dvd_neg_iff_dvd
theorem neg_dvd_iff_dvd : ∀a b : , (-a | b) ↔ (a | b) := algebra.neg_dvd_iff_dvd theorem neg_dvd_iff_dvd : ∀a b : , -a b ↔ a b := algebra.neg_dvd_iff_dvd
theorem dvd_sub : ∀a b c : , (a | b) → (a | c) → (a | b - c) := algebra.dvd_sub theorem dvd_sub : ∀a b c : , a b → a c → a b - c := algebra.dvd_sub
theorem mul_ne_zero : ∀{a b : }, a ≠ 0 → b ≠ 0 → a * b ≠ 0 := @algebra.mul_ne_zero _ _ theorem mul_ne_zero : ∀{a b : }, a ≠ 0 → b ≠ 0 → a * b ≠ 0 := @algebra.mul_ne_zero _ _
theorem mul.cancel_right : ∀{a b c : }, a ≠ 0 → b * a = c * a → b = c := theorem mul.cancel_right : ∀{a b c : }, a ≠ 0 → b * a = c * a → b = c :=
@algebra.mul.cancel_right _ _ @algebra.mul.cancel_right _ _
@ -769,9 +769,9 @@ section port_algebra
algebra.mul_self_eq_mul_self_iff algebra.mul_self_eq_mul_self_iff
theorem mul_self_eq_one_iff : ∀a : , a * a = 1 ↔ a = 1 a = -1 := theorem mul_self_eq_one_iff : ∀a : , a * a = 1 ↔ a = 1 a = -1 :=
algebra.mul_self_eq_one_iff algebra.mul_self_eq_one_iff
theorem dvd_of_mul_dvd_mul_left : ∀{a b c : }, a ≠ 0 → (a * b | a * c) → (b | c) := theorem dvd_of_mul_dvd_mul_left : ∀{a b c : }, a ≠ 0 → a*b a*c → b c :=
@algebra.dvd_of_mul_dvd_mul_left _ _ @algebra.dvd_of_mul_dvd_mul_left _ _
theorem dvd_of_mul_dvd_mul_right : ∀{a b c : }, a ≠ 0 → (b * a | c * a) → (b | c) := theorem dvd_of_mul_dvd_mul_right : ∀{a b c : }, a ≠ 0 → b*a c*a → b c :=
@algebra.dvd_of_mul_dvd_mul_right _ _ @algebra.dvd_of_mul_dvd_mul_right _ _
end port_algebra end port_algebra

View file

@ -535,8 +535,8 @@ section port_algebra
theorem sign_mul : ∀a b : , sign (a * b) = sign a * sign b := algebra.sign_mul theorem sign_mul : ∀a b : , sign (a * b) = sign a * sign b := algebra.sign_mul
theorem abs_eq_sign_mul : ∀a : , abs a = sign a * a := algebra.abs_eq_sign_mul theorem abs_eq_sign_mul : ∀a : , abs a = sign a * a := algebra.abs_eq_sign_mul
theorem eq_sign_mul_abs : ∀a : , a = sign a * abs a := algebra.eq_sign_mul_abs theorem eq_sign_mul_abs : ∀a : , a = sign a * abs a := algebra.eq_sign_mul_abs
theorem abs_dvd_iff_dvd : ∀a b : , (abs a | b) ↔ (a | b) := algebra.abs_dvd_iff_dvd theorem abs_dvd_iff_dvd : ∀a b : , abs a b ↔ a b := algebra.abs_dvd_iff_dvd
theorem dvd_abs_iff : ∀a b : , (a | abs b) ↔ (a | b) := algebra.dvd_abs_iff theorem dvd_abs_iff : ∀a b : , a abs b ↔ a b := algebra.dvd_abs_iff
theorem abs_mul : ∀a b : , abs (a * b) = abs a * abs b := algebra.abs_mul theorem abs_mul : ∀a b : , abs (a * b) = abs a * abs b := algebra.abs_mul
theorem abs_mul_self : ∀a : , abs a * abs a = a * a := algebra.abs_mul_self theorem abs_mul_self : ∀a : , abs a * abs a = a * a := algebra.abs_mul_self
end port_algebra end port_algebra

View file

@ -315,36 +315,36 @@ section port_algebra
theorem mul.right_comm : ∀a b c : , (a * b) * c = (a * c) * b := algebra.mul.right_comm theorem mul.right_comm : ∀a b c : , (a * b) * c = (a * c) * b := algebra.mul.right_comm
definition dvd (a b : ) : Prop := algebra.dvd a b definition dvd (a b : ) : Prop := algebra.dvd a b
notation (a | b) := dvd a b notation a b := dvd a b
theorem dvd.intro : ∀{a b c : } (H : a * c = b), (a | b) := @algebra.dvd.intro _ _ theorem dvd.intro : ∀{a b c : } (H : a * c = b), a b := @algebra.dvd.intro _ _
theorem dvd.intro_left : ∀{a b c : } (H : c * a = b), (a | b) := @algebra.dvd.intro_left _ _ theorem dvd.intro_left : ∀{a b c : } (H : c * a = b), a b := @algebra.dvd.intro_left _ _
theorem exists_eq_mul_right_of_dvd : ∀{a b : } (H : (a | b)), ∃c, b = a * c := theorem exists_eq_mul_right_of_dvd : ∀{a b : } (H : a b), ∃c, b = a * c :=
@algebra.exists_eq_mul_right_of_dvd _ _ @algebra.exists_eq_mul_right_of_dvd _ _
theorem dvd.elim : ∀{P : Prop} {a b : } (H₁ : (a | b)) (H₂ : ∀c, b = a * c → P), P := theorem dvd.elim : ∀{P : Prop} {a b : } (H₁ : a b) (H₂ : ∀c, b = a * c → P), P :=
@algebra.dvd.elim _ _ @algebra.dvd.elim _ _
theorem exists_eq_mul_left_of_dvd : ∀{a b : } (H : (a | b)), ∃c, b = c * a := theorem exists_eq_mul_left_of_dvd : ∀{a b : } (H : a b), ∃c, b = c * a :=
@algebra.exists_eq_mul_left_of_dvd _ _ @algebra.exists_eq_mul_left_of_dvd _ _
theorem dvd.elim_left : ∀{P : Prop} {a b : } (H₁ : (a | b)) (H₂ : ∀c, b = c * a → P), P := theorem dvd.elim_left : ∀{P : Prop} {a b : } (H₁ : a b) (H₂ : ∀c, b = c * a → P), P :=
@algebra.dvd.elim_left _ _ @algebra.dvd.elim_left _ _
theorem dvd.refl : ∀a : , (a | a) := algebra.dvd.refl theorem dvd.refl : ∀a : , a a := algebra.dvd.refl
theorem dvd.trans : ∀{a b c : } (H₁ : (a | b)) (H₂ : (b | c)), (a | c) := @algebra.dvd.trans _ _ theorem dvd.trans : ∀{a b c : }, a b → b c → a c := @algebra.dvd.trans _ _
theorem eq_zero_of_zero_dvd : ∀{a : } (H : (0 | a)), a = 0 := @algebra.eq_zero_of_zero_dvd _ _ theorem eq_zero_of_zero_dvd : ∀{a : }, 0 a → a = 0 := @algebra.eq_zero_of_zero_dvd _ _
theorem dvd_zero : ∀a : , (a | 0) := algebra.dvd_zero theorem dvd_zero : ∀a : , a 0 := algebra.dvd_zero
theorem one_dvd : ∀a : , (1 | a) := algebra.one_dvd theorem one_dvd : ∀a : , 1 a := algebra.one_dvd
theorem dvd_mul_right : ∀a b : , (a | a * b) := algebra.dvd_mul_right theorem dvd_mul_right : ∀a b : , a a * b := algebra.dvd_mul_right
theorem dvd_mul_left : ∀a b : , (a | b * a) := algebra.dvd_mul_left theorem dvd_mul_left : ∀a b : , a b * a := algebra.dvd_mul_left
theorem dvd_mul_of_dvd_left : ∀{a b : } (H : (a | b)) (c : ), (a | b * c) := theorem dvd_mul_of_dvd_left : ∀{a b : } (H : a b) (c : ), a b * c :=
@algebra.dvd_mul_of_dvd_left _ _ @algebra.dvd_mul_of_dvd_left _ _
theorem dvd_mul_of_dvd_right : ∀{a b : } (H : (a | b)) (c : ), (a | c * b) := theorem dvd_mul_of_dvd_right : ∀{a b : } (H : a b) (c : ), a c * b :=
@algebra.dvd_mul_of_dvd_right _ _ @algebra.dvd_mul_of_dvd_right _ _
theorem mul_dvd_mul : ∀{a b c d : }, (a | b) → (c | d) → (a * c | b * d) := theorem mul_dvd_mul : ∀{a b c d : }, a b → c d → a * c b * d :=
@algebra.mul_dvd_mul _ _ @algebra.mul_dvd_mul _ _
theorem dvd_of_mul_right_dvd : ∀{a b c : }, (a * b | c) → (a | c) := theorem dvd_of_mul_right_dvd : ∀{a b c : }, a * b c → a c :=
@algebra.dvd_of_mul_right_dvd _ _ @algebra.dvd_of_mul_right_dvd _ _
theorem dvd_of_mul_left_dvd : ∀{a b c : }, (a * b | c) → (b | c) := theorem dvd_of_mul_left_dvd : ∀{a b c : }, a * b c → b c :=
@algebra.dvd_of_mul_left_dvd _ _ @algebra.dvd_of_mul_left_dvd _ _
theorem dvd_add : ∀{a b c : }, (a | b) → (a | c) → (a | b + c) := @algebra.dvd_add _ _ theorem dvd_add : ∀{a b c : }, a b → a c → a b + c := @algebra.dvd_add _ _
end port_algebra end port_algebra
end nat end nat

View file

@ -332,28 +332,28 @@ calc
/- divides -/ /- divides -/
theorem dvd_of_mod_eq_zero {m n : } (H : n mod m = 0) : (m | n) := theorem dvd_of_mod_eq_zero {m n : } (H : n mod m = 0) : m n :=
dvd.intro (!mul.comm ▸ div_mul_cancel_of_mod_eq_zero H) dvd.intro (!mul.comm ▸ div_mul_cancel_of_mod_eq_zero H)
theorem mod_eq_zero_of_dvd {m n : } (H : (m | n)) : n mod m = 0 := theorem mod_eq_zero_of_dvd {m n : } (H : m n) : n mod m = 0 :=
dvd.elim H dvd.elim H
(take z, (take z,
assume H1 : n = m * z, assume H1 : n = m * z,
H1⁻¹ ▸ !mul_mod_right) H1⁻¹ ▸ !mul_mod_right)
theorem dvd_iff_mod_eq_zero (m n : ) : (m | n) ↔ n mod m = 0 := theorem dvd_iff_mod_eq_zero (m n : ) : m n ↔ n mod m = 0 :=
iff.intro mod_eq_zero_of_dvd dvd_of_mod_eq_zero iff.intro mod_eq_zero_of_dvd dvd_of_mod_eq_zero
definition dvd.decidable_rel [instance] : decidable_rel dvd := definition dvd.decidable_rel [instance] : decidable_rel dvd :=
take m n, decidable_of_decidable_of_iff _ (iff.symm !dvd_iff_mod_eq_zero) take m n, decidable_of_decidable_of_iff _ (iff.symm !dvd_iff_mod_eq_zero)
theorem div_mul_cancel {m n : } (H : (n | m)) : m div n * n = m := theorem div_mul_cancel {m n : } (H : n m) : m div n * n = m :=
div_mul_cancel_of_mod_eq_zero (mod_eq_zero_of_dvd H) div_mul_cancel_of_mod_eq_zero (mod_eq_zero_of_dvd H)
theorem mul_div_cancel' {m n : } (H : (n | m)) : n * (m div n) = m := theorem mul_div_cancel' {m n : } (H : n m) : n * (m div n) = m :=
!mul.comm ▸ div_mul_cancel H !mul.comm ▸ div_mul_cancel H
theorem eq_mul_of_div_eq {m n k : } (H1 : (m | n)) (H2 : n div m = k) : n = m * k := theorem eq_mul_of_div_eq {m n k : } (H1 : m n) (H2 : n div m = k) : n = m * k :=
eq.symm (calc eq.symm (calc
m * k = m * (n div m) : H2 m * k = m * (n div m) : H2
... = n : mul_div_cancel' H1) ... = n : mul_div_cancel' H1)
@ -363,7 +363,7 @@ calc
n = n * k div k : mul_div_cancel _ H1 n = n * k div k : mul_div_cancel _ H1
... = m div k : H2 ... = m div k : H2
theorem dvd_of_dvd_add_left {m n₁ n₂ : } (H₁ : (m | n₁ + n₂)) (H₂ : (m | n₁)) : (m | n₂) := theorem dvd_of_dvd_add_left {m n₁ n₂ : } (H₁ : m n₁ + n₂) (H₂ : m n₁) : m n₂ :=
obtain (c₁ : nat) (Hc₁ : n₁ + n₂ = m * c₁), from H₁, obtain (c₁ : nat) (Hc₁ : n₁ + n₂ = m * c₁), from H₁,
obtain (c₂ : nat) (Hc₂ : n₁ = m * c₂), from H₂, obtain (c₂ : nat) (Hc₂ : n₁ = m * c₂), from H₂,
have aux : m * (c₁ - c₂) = n₂, from calc have aux : m * (c₁ - c₂) = n₂, from calc
@ -373,25 +373,25 @@ have aux : m * (c₁ - c₂) = n₂, from calc
... = n₂ : add_sub_cancel_left, ... = n₂ : add_sub_cancel_left,
dvd.intro aux dvd.intro aux
theorem dvd_of_dvd_add_right {m n1 n2 : } (H : (m | (n1 + n2))) : (m | n2) → (m | n1) := theorem dvd_of_dvd_add_right {m n1 n2 : } (H : m (n1 + n2)) : m n2 → m n1 :=
dvd_of_dvd_add_left (!add.comm ▸ H) dvd_of_dvd_add_left (!add.comm ▸ H)
theorem dvd_sub {m n1 n2 : } (H1 : (m | n1)) (H2 : (m | n2)) : (m | n1 - n2) := theorem dvd_sub {m n1 n2 : } (H1 : m n1) (H2 : m n2) : m n1 - n2 :=
by_cases by_cases
(assume H3 : n1 ≥ n2, (assume H3 : n1 ≥ n2,
have H4 : n1 = n1 - n2 + n2, from (sub_add_cancel H3)⁻¹, have H4 : n1 = n1 - n2 + n2, from (sub_add_cancel H3)⁻¹,
show (m | n1 - n2), from dvd_of_dvd_add_right (H4 ▸ H1) H2) show m n1 - n2, from dvd_of_dvd_add_right (H4 ▸ H1) H2)
(assume H3 : ¬ (n1 ≥ n2), (assume H3 : ¬ (n1 ≥ n2),
have H4 : n1 - n2 = 0, from sub_eq_zero_of_le (le_of_lt (lt_of_not_le H3)), have H4 : n1 - n2 = 0, from sub_eq_zero_of_le (le_of_lt (lt_of_not_le H3)),
show (m | n1 - n2), from H4⁻¹ ▸ dvd_zero _) show m n1 - n2, from H4⁻¹ ▸ dvd_zero _)
theorem dvd.antisymm {m n : } : (m | n) → (n | m) → m = n := theorem dvd.antisymm {m n : } : m n → n m → m = n :=
by_cases_zero_pos n by_cases_zero_pos n
(assume H1, assume H2 : (0 | m), eq_zero_of_zero_dvd H2) (assume H1, assume H2 : 0 m, eq_zero_of_zero_dvd H2)
(take n, (take n,
assume Hpos : n > 0, assume Hpos : n > 0,
assume H1 : (m | n), assume H1 : m n,
assume H2 : (n | m), assume H2 : n m,
obtain k (Hk : n = m * k), from exists_eq_mul_right_of_dvd H1, obtain k (Hk : n = m * k), from exists_eq_mul_right_of_dvd H1,
obtain l (Hl : m = n * l), from exists_eq_mul_right_of_dvd H2, obtain l (Hl : m = n * l), from exists_eq_mul_right_of_dvd H2,
have H3 : n * (l * k) = n, from !mul.assoc ▸ Hl ▸ Hk⁻¹, have H3 : n * (l * k) = n, from !mul.assoc ▸ Hl ▸ Hk⁻¹,
@ -399,7 +399,7 @@ by_cases_zero_pos n
have H5 : k = 1, from eq_one_of_mul_eq_one_left H4, have H5 : k = 1, from eq_one_of_mul_eq_one_left H4,
show m = n, from (mul_one m)⁻¹ ⬝ (H5 ▸ Hk⁻¹)) show m = n, from (mul_one m)⁻¹ ⬝ (H5 ▸ Hk⁻¹))
theorem mul_div_assoc (m : ) {n k : } (H : (k | n)) : m * n div k = m * (n div k) := theorem mul_div_assoc (m : ) {n k : } (H : k n) : m * n div k = m * (n div k) :=
or.elim (eq_zero_or_pos k) or.elim (eq_zero_or_pos k)
(assume H1 : k = 0, (assume H1 : k = 0,
calc calc
@ -415,17 +415,17 @@ or.elim (eq_zero_or_pos k)
... = m * (n div k) * k div k : mul.assoc ... = m * (n div k) * k div k : mul.assoc
... = m * (n div k) : mul_div_cancel _ H1) ... = m * (n div k) : mul_div_cancel _ H1)
theorem dvd_of_mul_dvd_mul_left {m n k : } (kpos : k > 0) (H : (k * m | k * n)) : (m | n) := theorem dvd_of_mul_dvd_mul_left {m n k : } (kpos : k > 0) (H : k * m k * n) : m n :=
dvd.elim H dvd.elim H
(take l, (take l,
assume H1 : k * n = k * m * l, assume H1 : k * n = k * m * l,
have H2 : n = m * l, from eq_of_mul_eq_mul_left kpos (H1 ⬝ !mul.assoc), have H2 : n = m * l, from eq_of_mul_eq_mul_left kpos (H1 ⬝ !mul.assoc),
dvd.intro H2⁻¹) dvd.intro H2⁻¹)
theorem dvd_of_mul_dvd_mul_right {m n k : } (kpos : k > 0) (H : (m * k | n * k)) : (m | n) := theorem dvd_of_mul_dvd_mul_right {m n k : } (kpos : k > 0) (H : m * k n * k) : m n :=
dvd_of_mul_dvd_mul_left kpos (!mul.comm ▸ !mul.comm ▸ H) dvd_of_mul_dvd_mul_left kpos (!mul.comm ▸ !mul.comm ▸ H)
theorem div_dvd_div {k m n : } (H1 : (k | m)) (H2 : (m | n)) : (m div k | n div k) := theorem div_dvd_div {k m n : } (H1 : k m) (H2 : m n) : m div k n div k :=
have H3 : m = m div k * k, from (div_mul_cancel H1)⁻¹, have H3 : m = m div k * k, from (div_mul_cancel H1)⁻¹,
have H4 : n = n div k * k, from (div_mul_cancel (dvd.trans H1 H2))⁻¹, have H4 : n = n div k * k, from (div_mul_cancel (dvd.trans H1 H2))⁻¹,
or.elim (eq_zero_or_pos k) or.elim (eq_zero_or_pos k)
@ -518,36 +518,36 @@ have aux : Q (m, n), from
H1 m (succ n₁) hlt₁ hp))), H1 m (succ n₁) hlt₁ hp))),
aux aux
theorem gcd_dvd (m n : ) : (gcd m n | m) ∧ (gcd m n | n) := theorem gcd_dvd (m n : ) : (gcd m n m) ∧ (gcd m n n) :=
gcd.induction m n gcd.induction m n
(take m, (take m,
show (gcd m 0 | m) ∧ (gcd m 0 | 0), by simp) show (gcd m 0 m) ∧ (gcd m 0 0), by simp)
(take m n, (take m n,
assume npos : 0 < n, assume npos : 0 < n,
assume IH : (gcd n (m mod n) | n) ∧ (gcd n (m mod n) | (m mod n)), assume IH : (gcd n (m mod n) n) ∧ (gcd n (m mod n) (m mod n)),
have H : (gcd n (m mod n) | (m div n * n + m mod n)), from have H : (gcd n (m mod n) (m div n * n + m mod n)), from
dvd_add (dvd.trans (and.elim_left IH) !dvd_mul_left) (and.elim_right IH), dvd_add (dvd.trans (and.elim_left IH) !dvd_mul_left) (and.elim_right IH),
have H1 : (gcd n (m mod n) | m), from eq_div_mul_add_mod⁻¹ ▸ H, have H1 : (gcd n (m mod n) m), from eq_div_mul_add_mod⁻¹ ▸ H,
have gcd_eq : gcd n (m mod n) = gcd m n, from !gcd_rec⁻¹, have gcd_eq : gcd n (m mod n) = gcd m n, from !gcd_rec⁻¹,
show (gcd m n | m) ∧ (gcd m n | n), from gcd_eq ▸ (and.intro H1 (and.elim_left IH))) show (gcd m n m) ∧ (gcd m n n), from gcd_eq ▸ (and.intro H1 (and.elim_left IH)))
theorem gcd_dvd_left (m n : ) : (gcd m n | m) := and.elim_left !gcd_dvd theorem gcd_dvd_left (m n : ) : gcd m n m := and.elim_left !gcd_dvd
theorem gcd_dvd_right (m n : ) : (gcd m n | n) := and.elim_right !gcd_dvd theorem gcd_dvd_right (m n : ) : gcd m n n := and.elim_right !gcd_dvd
theorem dvd_gcd {m n k : } : (k | m) → (k | n) → (k | gcd m n) := theorem dvd_gcd {m n k : } : k m → k n → k gcd m n :=
gcd.induction m n gcd.induction m n
(take m, assume (h₁ : (k | m)) (h₂ : (k | 0)), (take m, assume (h₁ : k m) (h₂ : k 0),
show (k | gcd m 0), from !gcd_zero_right⁻¹ ▸ h₁) show k gcd m 0, from !gcd_zero_right⁻¹ ▸ h₁)
(take m n, (take m n,
assume npos : n > 0, assume npos : n > 0,
assume IH : (k | n) → (k | m mod n) → (k | gcd n (m mod n)), assume IH : k n → k m mod n → k gcd n (m mod n),
assume H1 : (k | m), assume H1 : k m,
assume H2 : (k | n), assume H2 : k n,
have H3 : (k | m div n * n + m mod n), from eq_div_mul_add_mod ▸ H1, have H3 : k m div n * n + m mod n, from eq_div_mul_add_mod ▸ H1,
have H4 : (k | m mod n), from nat.dvd_of_dvd_add_left H3 (dvd.trans H2 (by simp)), have H4 : k m mod n, from nat.dvd_of_dvd_add_left H3 (dvd.trans H2 (by simp)),
have gcd_eq : gcd n (m mod n) = gcd m n, from !gcd_rec⁻¹, have gcd_eq : gcd n (m mod n) = gcd m n, from !gcd_rec⁻¹,
show (k | gcd m n), from gcd_eq ▸ IH H2 H4) show k gcd m n, from gcd_eq ▸ IH H2 H4)
theorem gcd.comm (m n : ) : gcd m n = gcd n m := theorem gcd.comm (m n : ) : gcd m n = gcd n m :=
dvd.antisymm dvd.antisymm
@ -603,7 +603,7 @@ or.elim (eq_zero_or_pos m)
theorem eq_zero_of_gcd_eq_zero_right {m n : } (H : gcd m n = 0) : n = 0 := theorem eq_zero_of_gcd_eq_zero_right {m n : } (H : gcd m n = 0) : n = 0 :=
eq_zero_of_gcd_eq_zero_left (!gcd.comm ▸ H) eq_zero_of_gcd_eq_zero_left (!gcd.comm ▸ H)
theorem gcd_div {m n k : } (H1 : (k | m)) (H2 : (k | n)) : gcd (m div k) (n div k) = gcd m n div k := theorem gcd_div {m n k : } (H1 : (k m)) (H2 : (k n)) : gcd (m div k) (n div k) = gcd m n div k :=
or.elim (eq_zero_or_pos k) or.elim (eq_zero_or_pos k)
(assume H3 : k = 0, (assume H3 : k = 0,
calc calc
@ -621,16 +621,16 @@ or.elim (eq_zero_or_pos k)
... = gcd m (n div k * k) : div_mul_cancel H1 ... = gcd m (n div k * k) : div_mul_cancel H1
... = gcd m n : div_mul_cancel H2)) ... = gcd m n : div_mul_cancel H2))
theorem gcd_dvd_gcd_mul_left (m n k : ) : (gcd m n | gcd (k * m) n) := theorem gcd_dvd_gcd_mul_left (m n k : ) : gcd m n gcd (k * m) n :=
dvd_gcd (dvd.trans !gcd_dvd_left !dvd_mul_left) !gcd_dvd_right dvd_gcd (dvd.trans !gcd_dvd_left !dvd_mul_left) !gcd_dvd_right
theorem gcd_dvd_gcd_mul_right (m n k : ) : (gcd m n | gcd (m * k) n) := theorem gcd_dvd_gcd_mul_right (m n k : ) : gcd m n gcd (m * k) n :=
!mul.comm ▸ !gcd_dvd_gcd_mul_left !mul.comm ▸ !gcd_dvd_gcd_mul_left
theorem gcd_dvd_gcd_mul_left_right (m n k : ) : (gcd m n | gcd m (k * n)) := theorem gcd_dvd_gcd_mul_left_right (m n k : ) : gcd m n gcd m (k * n) :=
dvd_gcd !gcd_dvd_left (dvd.trans !gcd_dvd_right !dvd_mul_left) dvd_gcd !gcd_dvd_left (dvd.trans !gcd_dvd_right !dvd_mul_left)
theorem gcd_dvd_gcd_mul_right_right (m n k : ) : (gcd m n | gcd m (n * k)) := theorem gcd_dvd_gcd_mul_right_right (m n k : ) : gcd m n gcd m (n * k) :=
!mul.comm ▸ !gcd_dvd_gcd_mul_left_right !mul.comm ▸ !gcd_dvd_gcd_mul_left_right
/- lcm -/ /- lcm -/
@ -669,17 +669,17 @@ calc
... = m * m div m : gcd_self ... = m * m div m : gcd_self
... = m : H ... = m : H
theorem dvd_lcm_left (m n : ) : (m | lcm m n) := theorem dvd_lcm_left (m n : ) : m lcm m n :=
have H : lcm m n = m * (n div gcd m n), from mul_div_assoc _ !gcd_dvd_right, have H : lcm m n = m * (n div gcd m n), from mul_div_assoc _ !gcd_dvd_right,
dvd.intro H⁻¹ dvd.intro H⁻¹
theorem dvd_lcm_right (m n : ) : (n | lcm m n) := theorem dvd_lcm_right (m n : ) : n lcm m n :=
!lcm.comm ▸ !dvd_lcm_left !lcm.comm ▸ !dvd_lcm_left
theorem gcd_mul_lcm (m n : ) : gcd m n * lcm m n = m * n := theorem gcd_mul_lcm (m n : ) : gcd m n * lcm m n = m * n :=
eq.symm (eq_mul_of_div_eq (dvd.trans !gcd_dvd_left !dvd_mul_right) rfl) eq.symm (eq_mul_of_div_eq (dvd.trans !gcd_dvd_left !dvd_mul_right) rfl)
theorem lcm_dvd {m n k : } (H1 : (m | k)) (H2 : (n | k)) : (lcm m n | k) := theorem lcm_dvd {m n k : } (H1 : m k) (H2 : n k) : lcm m n k :=
or.elim (eq_zero_or_pos k) or.elim (eq_zero_or_pos k)
(assume kzero : k = 0, !kzero⁻¹ ▸ !dvd_zero) (assume kzero : k = 0, !kzero⁻¹ ▸ !dvd_zero)
(assume kpos : k > 0, (assume kpos : k > 0,
@ -736,16 +736,16 @@ definition coprime [reducible] (m n : ) : Prop := gcd m n = 1
theorem coprime_swap {m n : } (H : coprime n m) : coprime m n := theorem coprime_swap {m n : } (H : coprime n m) : coprime m n :=
!gcd.comm ▸ H !gcd.comm ▸ H
theorem dvd_of_coprime_of_dvd_mul_right {m n k : } (H1 : coprime k n) (H2 : (k | m * n)) : (k | m) := theorem dvd_of_coprime_of_dvd_mul_right {m n k : } (H1 : coprime k n) (H2 : k m * n) : k m :=
have H3 : gcd (m * k) (m * n) = m, from have H3 : gcd (m * k) (m * n) = m, from
calc calc
gcd (m * k) (m * n) = m * gcd k n : gcd_mul_left gcd (m * k) (m * n) = m * gcd k n : gcd_mul_left
... = m * 1 : H1 ... = m * 1 : H1
... = m : mul_one, ... = m : mul_one,
have H4 : (k | gcd (m * k) (m * n)), from dvd_gcd !dvd_mul_left H2, have H4 : (k gcd (m * k) (m * n)), from dvd_gcd !dvd_mul_left H2,
H3 ▸ H4 H3 ▸ H4
theorem dvd_of_coprime_of_dvd_mul_left {m n k : } (H1 : coprime k m) (H2 : (k | m * n)) : (k | n) := theorem dvd_of_coprime_of_dvd_mul_left {m n k : } (H1 : coprime k m) (H2 : k m * n) : k n :=
dvd_of_coprime_of_dvd_mul_right H1 (!mul.comm ▸ H2) dvd_of_coprime_of_dvd_mul_right H1 (!mul.comm ▸ H2)
theorem gcd_mul_left_cancel_of_coprime {k : } (m : ) {n : } (H : coprime k n) : theorem gcd_mul_left_cancel_of_coprime {k : } (m : ) {n : } (H : coprime k n) :
@ -796,7 +796,7 @@ theorem coprime_mul_right {k m n : } (H1 : coprime k m) (H2 : coprime k n) :
coprime_swap (coprime_mul (coprime_swap H1) (coprime_swap H2)) coprime_swap (coprime_mul (coprime_swap H1) (coprime_swap H2))
theorem coprime_of_coprime_mul_left {k m n : } (H : coprime (k * m) n) : coprime m n := theorem coprime_of_coprime_mul_left {k m n : } (H : coprime (k * m) n) : coprime m n :=
have H1 : (gcd m n | gcd (k * m) n), from !gcd_dvd_gcd_mul_left, have H1 : (gcd m n gcd (k * m) n), from !gcd_dvd_gcd_mul_left,
eq_one_of_dvd_one (H ▸ H1) eq_one_of_dvd_one (H ▸ H1)
theorem coprime_of_coprime_mul_right {k m n : } (H : coprime (m * k) n) : coprime m n := theorem coprime_of_coprime_mul_right {k m n : } (H : coprime (m * k) n) : coprime m n :=
@ -808,27 +808,27 @@ coprime_swap (coprime_of_coprime_mul_left (coprime_swap H))
theorem coprime_of_coprime_mul_right_right {k m n : } (H : coprime m (n * k)) : coprime m n := theorem coprime_of_coprime_mul_right_right {k m n : } (H : coprime m (n * k)) : coprime m n :=
coprime_of_coprime_mul_left_right (!mul.comm ▸ H) coprime_of_coprime_mul_left_right (!mul.comm ▸ H)
theorem exists_eq_prod_and_dvd_and_dvd {m n k} (H : (k | m * n)) : theorem exists_eq_prod_and_dvd_and_dvd {m n k} (H : k m * n) :
∃ m' n', k = m' * n' ∧ (m' | m) ∧ (n' | n) := ∃ m' n', k = m' * n' ∧ m' m ∧ n' n :=
or.elim (eq_zero_or_pos (gcd k m)) or.elim (eq_zero_or_pos (gcd k m))
(assume H1 : gcd k m = 0, (assume H1 : gcd k m = 0,
have H2 : k = 0, from eq_zero_of_gcd_eq_zero_left H1, have H2 : k = 0, from eq_zero_of_gcd_eq_zero_left H1,
have H3 : m = 0, from eq_zero_of_gcd_eq_zero_right H1, have H3 : m = 0, from eq_zero_of_gcd_eq_zero_right H1,
have H4 : k = 0 * n, from H2 ⬝ !zero_mul⁻¹, have H4 : k = 0 * n, from H2 ⬝ !zero_mul⁻¹,
have H5 : (0 | m), from H3⁻¹ ▸ !dvd.refl, have H5 : 0 m, from H3⁻¹ ▸ !dvd.refl,
have H6 : (n | n), from !dvd.refl, have H6 : n n, from !dvd.refl,
exists.intro _ (exists.intro _ (and.intro H4 (and.intro H5 H6)))) exists.intro _ (exists.intro _ (and.intro H4 (and.intro H5 H6))))
(assume H1 : gcd k m > 0, (assume H1 : gcd k m > 0,
have H2 : (gcd k m | k), from !gcd_dvd_left, have H2 : gcd k m k, from !gcd_dvd_left,
have H3 : (k div gcd k m | (m * n) div gcd k m), from div_dvd_div H2 H, have H3 : k div gcd k m (m * n) div gcd k m, from div_dvd_div H2 H,
have H4 : (m * n) div gcd k m = (m div gcd k m) * n, from have H4 : (m * n) div gcd k m = (m div gcd k m) * n, from
calc calc
m * n div gcd k m = n * m div gcd k m : mul.comm m * n div gcd k m = n * m div gcd k m : mul.comm
... = n * (m div gcd k m) : !mul_div_assoc !gcd_dvd_right ... = n * (m div gcd k m) : !mul_div_assoc !gcd_dvd_right
... = m div gcd k m * n : mul.comm, ... = m div gcd k m * n : mul.comm,
have H5 : (k div gcd k m | (m div gcd k m) * n), from H4 ▸ H3, have H5 : k div gcd k m (m div gcd k m) * n, from H4 ▸ H3,
have H6 : coprime (k div gcd k m) (m div gcd k m), from coprime_div_gcd_div_gcd H1, have H6 : coprime (k div gcd k m) (m div gcd k m), from coprime_div_gcd_div_gcd H1,
have H7 : (k div gcd k m | n), from dvd_of_coprime_of_dvd_mul_left H6 H5, have H7 : k div gcd k m n, from dvd_of_coprime_of_dvd_mul_left H6 H5,
have H8 : k = gcd k m * (k div gcd k m), from (mul_div_cancel' H2)⁻¹, have H8 : k = gcd k m * (k div gcd k m), from (mul_div_cancel' H2)⁻¹,
exists.intro _ (exists.intro _ (and.intro H8 (and.intro !gcd_dvd_right H7)))) exists.intro _ (exists.intro _ (and.intro H8 (and.intro !gcd_dvd_right H7))))

View file

@ -353,7 +353,7 @@ ne.symm (ne_of_lt H)
theorem exists_eq_succ_of_pos {n : } (H : n > 0) : exists l, n = succ l := theorem exists_eq_succ_of_pos {n : } (H : n > 0) : exists l, n = succ l :=
exists_eq_succ_of_lt H exists_eq_succ_of_lt H
theorem pos_of_dvd_of_pos {m n : } (H1 : (m | n)) (H2 : n > 0) : m > 0 := theorem pos_of_dvd_of_pos {m n : } (H1 : m n) (H2 : n > 0) : m > 0 :=
pos_of_ne_zero pos_of_ne_zero
(assume H3 : m = 0, (assume H3 : m = 0,
have H4 : n = 0, from eq_zero_of_zero_dvd (H3 ▸ H1), have H4 : n = 0, from eq_zero_of_zero_dvd (H3 ▸ H1),
@ -412,7 +412,7 @@ eq_of_mul_eq_mul_right Hpos (H ⬝ !one_mul⁻¹)
theorem eq_one_of_mul_eq_self_right {n m : } (Hpos : m > 0) (H : m * n = m) : n = 1 := theorem eq_one_of_mul_eq_self_right {n m : } (Hpos : m > 0) (H : m * n = m) : n = 1 :=
eq_one_of_mul_eq_self_left Hpos (!mul.comm ▸ H) eq_one_of_mul_eq_self_left Hpos (!mul.comm ▸ H)
theorem eq_one_of_dvd_one {n : } (H : (n | 1)) : n = 1 := theorem eq_one_of_dvd_one {n : } (H : n 1) : n = 1 :=
dvd.elim H dvd.elim H
(take m, (take m,
assume H1 : 1 = n * m, assume H1 : 1 = n * m,

View file

@ -89,7 +89,7 @@ reserve infixl ``:65
/- other symbols -/ /- other symbols -/
reserve notation `(` a `|` b `)` reserve infix ``:50
reserve infixl `++`:65 reserve infixl `++`:65
reserve infixr `::`:65 reserve infixr `::`:65

View file

@ -92,13 +92,13 @@ opaque definition change (e : expr) : tactic := builtin
opaque definition assert_hypothesis (id : expr) (e : expr) : tactic := builtin opaque definition assert_hypothesis (id : expr) (e : expr) : tactic := builtin
infixl `;`:15 := and_then infixl `;`:15 := and_then
notation `[` h `|` r:(foldl `|` (e r, or_else r e) h) `]` := r notation `(` h `|` r:(foldl `|` (e r, or_else r e) h) `)` := r
definition try (t : tactic) : tactic := [t | id] definition try (t : tactic) : tactic := (t | id)
definition repeat1 (t : tactic) : tactic := t ; repeat t definition repeat1 (t : tactic) : tactic := t ; repeat t
definition focus (t : tactic) : tactic := focus_at t 0 definition focus (t : tactic) : tactic := focus_at t 0
definition determ (t : tactic) : tactic := at_most t 1 definition determ (t : tactic) : tactic := at_most t 1
definition trivial : tactic := [ apply eq.refl | apply true.intro | assumption ] definition trivial : tactic := (apply eq.refl | apply true.intro | assumption)
definition do (n : num) (t : tactic) : tactic := definition do (n : num) (t : tactic) : tactic :=
nat.rec id (λn t', (t;t')) (nat.of_num n) nat.rec id (λn t', (t;t')) (nat.of_num n)

View file

@ -1403,7 +1403,11 @@ expr parser::parse_tactic_nud() {
} else { } else {
return parse_expr(); return parse_expr();
} }
} else if (curr_is_token(get_lbracket_tk())) { } else if (curr_is_token_or_id(get_rewrite_tk())) {
auto p = pos();
next();
return save_pos(parse_rewrite_tactic(*this), p);
} else if (curr_is_token(get_lparen_tk())) {
next(); next();
expr r = parse_tactic(); expr r = parse_tactic();
while (curr_is_token(get_bar_tk())) { while (curr_is_token(get_bar_tk())) {
@ -1412,15 +1416,6 @@ expr parser::parse_tactic_nud() {
expr n = parse_tactic(); expr n = parse_tactic();
r = mk_app({save_pos(mk_constant(get_tactic_or_else_name()), bar_pos), r, n}, bar_pos); r = mk_app({save_pos(mk_constant(get_tactic_or_else_name()), bar_pos), r, n}, bar_pos);
} }
check_token_next(get_rbracket_tk(), "invalid or-else tactic, ']' expected");
return r;
} else if (curr_is_token_or_id(get_rewrite_tk())) {
auto p = pos();
next();
return save_pos(parse_rewrite_tactic(*this), p);
} else if (curr_is_token(get_lparen_tk())) {
next();
expr r = parse_tactic();
check_token_next(get_rparen_tk(), "invalid tactic, ')' expected"); check_token_next(get_rparen_tk(), "invalid tactic, ')' expected");
return r; return r;
} else { } else {

View file

@ -1,7 +1,7 @@
import data.nat import data.nat
open nat eq.ops open nat eq.ops
theorem lcm_dvd {m n k : nat} (H1 : (m | k)) (H2 : (n | k)) : (lcm m n | k) := theorem lcm_dvd {m n k : nat} (H1 : m k) (H2 : (n k)) : (lcm m n k) :=
match eq_zero_or_pos k with match eq_zero_or_pos k with
| @or.inl _ _ kzero := | @or.inl _ _ kzero :=
begin begin

View file

@ -15,6 +15,6 @@ theorem inh_fun [instance] {A B : Type} [H : inh B] : inh (A → B)
theorem tst {A B : Type} (H : inh B) : inh (A → B → B) theorem tst {A B : Type} (H : inh B) : inh (A → B → B)
theorem T1 {A : Type} (a : A) : inh A := theorem T1 {A : Type} (a : A) : inh A :=
by repeat [apply @inh.intro | eassumption] by repeat (apply @inh.intro | eassumption)
theorem T2 : inh Prop theorem T2 : inh Prop

View file

@ -4,5 +4,5 @@ open nat
example : is_true (2 = 2) := trivial example : is_true (2 = 2) := trivial
example : is_false (3 = 2) := trivial example : is_false (3 = 2) := trivial
example : is_true (2 < 3) := trivial example : is_true (2 < 3) := trivial
example : is_true (3 | 9) := trivial example : is_true (3 9) := trivial
example : is_false (3 | 7) := trivial example : is_false (3 7) := trivial

View file

@ -16,7 +16,7 @@ begin
match H with match H with
| ⟪ H₁, H₂, H₃, H₄ ⟫ := | ⟪ H₁, H₂, H₃, H₄ ⟫ :=
begin begin
repeat [apply and.intro | assumption] repeat (apply and.intro | assumption)
end end
end end
end end
@ -27,7 +27,7 @@ print definition tst
theorem tst2 (a b c d : Prop) : a ∧ b ∧ c ∧ d ↔ d ∧ c ∧ b ∧ a := theorem tst2 (a b c d : Prop) : a ∧ b ∧ c ∧ d ↔ d ∧ c ∧ b ∧ a :=
begin begin
apply iff.intro, apply iff.intro,
repeat (intro H; repeat [cases H with [H', H] | apply and.intro | assumption]) repeat (intro H; repeat (cases H with [H', H] | apply and.intro | assumption))
end end
print definition tst2 print definition tst2

View file

@ -2,7 +2,7 @@ import logic
open tactic open tactic
definition basic_tac : tactic definition basic_tac : tactic
:= repeat [apply @and.intro | assumption] := repeat (apply @and.intro | assumption)
set_begin_end_tactic basic_tac -- basic_tac is automatically applied to each element of a proof-qed block set_begin_end_tactic basic_tac -- basic_tac is automatically applied to each element of a proof-qed block

View file

@ -2,4 +2,4 @@ import logic
open tactic open tactic
theorem T (a b c d : Prop) (Ha : a) (Hb : b) (Hc : c) (Hd : d) : a ∧ b ∧ c ∧ d theorem T (a b c d : Prop) (Ha : a) (Hb : b) (Hc : c) (Hd : d) : a ∧ b ∧ c ∧ d
:= by fixpoint (λ f, [apply @and.intro; f | assumption; f | id]) := by fixpoint (λ f, (apply @and.intro; f | assumption; f | id))

View file

@ -13,9 +13,9 @@ theorem T1 {A : Type.{2}} (a : A) : a = a
theorem T2 {a b c : Prop} (Ha : a) (Hb : b) (Hc : c) : a ∧ b ∧ c theorem T2 {a b c : Prop} (Ha : a) (Hb : b) (Hc : c) : a ∧ b ∧ c
:= _ := _
definition my_tac3 := fixpoint (λ f, [apply @or.intro_left; f | definition my_tac3 := fixpoint (λ f, (apply @or.intro_left; f |
apply @or.intro_right; f | apply @or.intro_right; f |
assumption]) assumption))
tactic_hint my_tac3 tactic_hint my_tac3

View file

@ -11,9 +11,9 @@ theorem T1 {A : Type.{2}} (a : A) : a = a
theorem T2 {a b c : Prop} (Ha : a) (Hb : b) (Hc : c) : a ∧ b ∧ c theorem T2 {a b c : Prop} (Ha : a) (Hb : b) (Hc : c) : a ∧ b ∧ c
definition my_tac3 := fixpoint (λ f, [apply @or.intro_left; f | definition my_tac3 := fixpoint (λ f, (apply @or.intro_left; f |
apply @or.intro_right; f | apply @or.intro_right; f |
assumption]) assumption))
tactic_hint my_tac3 tactic_hint my_tac3
theorem T3 {a b c : Prop} (Hb : b) : a b c theorem T3 {a b c : Prop} (Hb : b) : a b c

View file

@ -12,10 +12,10 @@ theorem inl_inhabited {A : Type} (B : Type) (H : inhabited A) : inhabited (sum A
theorem inr_inhabited (A : Type) {B : Type} (H : inhabited B) : inhabited (sum A B) theorem inr_inhabited (A : Type) {B : Type} (H : inhabited B) : inhabited (sum A B)
:= inhabited.destruct H (λ b, inhabited.mk (sum.inr A b)) := inhabited.destruct H (λ b, inhabited.mk (sum.inr A b))
definition my_tac := fixpoint (λ t, [ apply @inl_inhabited; t definition my_tac := fixpoint (λ t, ( apply @inl_inhabited; t
| apply @inr_inhabited; t | apply @inr_inhabited; t
| apply @num.is_inhabited | apply @num.is_inhabited
]) ))
tactic_hint my_tac tactic_hint my_tac

View file

@ -1,9 +1,7 @@
import logic import logic
open tactic open tactic
definition my_tac := repeat ([ apply @and.intro definition my_tac := repeat (apply @and.intro | apply @eq.refl)
| apply @eq.refl
])
tactic_hint my_tac tactic_hint my_tac
theorem T1 {A : Type} {B : Type} (a : A) (b c : B) : a = a ∧ b = b ∧ c = c theorem T1 {A : Type} {B : Type} (a : A) (b c : B) : a = a ∧ b = b ∧ c = c

View file

@ -2,8 +2,8 @@ import logic
open tactic open tactic
theorem tst {A B : Prop} (H1 : A) (H2 : B) : A theorem tst {A B : Prop} (H1 : A) (H2 : B) : A
:= by [trace "first"; state; now | := by (trace "first"; state; now |
trace "second"; state; fail | trace "second"; state; fail |
trace "third"; assumption] trace "third"; assumption)
check tst check tst

View file

@ -6,6 +6,6 @@ theorem tst {A B : Prop} (H1 : A) (H2 : B) : A ∧ B ∧ A
check tst check tst
theorem tst2 {A B : Prop} (H1 : A) (H2 : B) : A ∧ B ∧ A theorem tst2 {A B : Prop} (H1 : A) (H2 : B) : A ∧ B ∧ A
:= by repeat ([apply @and.intro | assumption] ; trace "STEP"; state; trace "----------") := by repeat ((apply @and.intro | assumption) ; trace "STEP"; state; trace "----------")
check tst2 check tst2