diff --git a/hott/algebra/precategory/functor.hlean b/hott/algebra/precategory/functor.hlean index 147698ccb..6dc90b47c 100644 --- a/hott/algebra/precategory/functor.hlean +++ b/hott/algebra/precategory/functor.hlean @@ -93,7 +93,7 @@ namespace functor fapply @is_iso.mk, apply (F (f⁻¹)), repeat (apply concat ; apply inverse ; apply (respect_comp F) ; 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) ), end diff --git a/hott/init/reserved_notation.hlean b/hott/init/reserved_notation.hlean index a8f96c8f5..ed6198e23 100644 --- a/hott/init/reserved_notation.hlean +++ b/hott/init/reserved_notation.hlean @@ -92,7 +92,7 @@ reserve infixl `∪`:65 /- other symbols -/ -reserve notation `(` a `|` b `)` +reserve infix `∣`:50 reserve infixl `++`:65 reserve infixr `::`:65 diff --git a/hott/init/tactic.hlean b/hott/init/tactic.hlean index 508574ded..054e0e33b 100644 --- a/hott/init/tactic.hlean +++ b/hott/init/tactic.hlean @@ -92,13 +92,13 @@ opaque definition change (e : expr) : tactic := builtin opaque definition assert_hypothesis (id : expr) (e : expr) : tactic := builtin 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 focus (t : tactic) : tactic := focus_at t 0 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 := nat.rec id (λn t', (t;t')) (nat.of_num n) diff --git a/library/algebra/ordered_ring.lean b/library/algebra/ordered_ring.lean index b1093d8bf..dd9d5e91e 100644 --- a/library/algebra/ordered_ring.lean +++ b/library/algebra/ordered_ring.lean @@ -544,10 +544,10 @@ section ... = -1 * abs a : by rewrite neg_eq_neg_one_mul ... = 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 - 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 theorem abs_mul (a b : A) : abs (a * b) = abs a * abs b := diff --git a/library/algebra/ring.lean b/library/algebra/ring.lean index 84f050724..67f32f299 100644 --- a/library/algebra/ring.lean +++ b/library/algebra/ring.lean @@ -72,28 +72,28 @@ section comm_semiring include s 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⁻¹ - 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) - 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₂ - 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)) - 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₃) - 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₁ (take d, assume H₃ : b = a * d, dvd.elim H₂ @@ -101,28 +101,28 @@ section comm_semiring dvd.intro (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) - 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 (take d, assume H₁ : b = a * d, dvd.intro (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 _) - 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 (take e, assume Haeb : b = a * e, dvd.elim dvd_cd @@ -131,13 +131,13 @@ section comm_semiring (show a * c * (e * f) = b * d, 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⁻¹)) - 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) - 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 (take d, assume Hadb : b = a * d, dvd.elim Hac @@ -259,35 +259,35 @@ section theorem mul_self_sub_one_eq : a * a - 1 = (a + 1) * (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 - (assume H : (a | -b), + (assume H : (a ∣ -b), dvd.elim H (take c, assume H' : -b = a * c, dvd.intro (show a * -c = b, by rewrite [-neg_mul_eq_mul_neg, -H', neg_neg]))) - (assume H : (a | b), + (assume H : (a ∣ b), dvd.elim H (take c, assume H' : b = a * c, dvd.intro (show a * -c = -b, 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 - (assume H : (-a | b), + (assume H : (-a ∣ b), dvd.elim H (take c, assume H' : b = -a * c, dvd.intro (show a * -c = b, by rewrite [-neg_mul_comm, H']))) - (assume H : (a | b), + (assume H : (a ∣ b), dvd.elim H (take c, assume H' : b = a * c, dvd.intro (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₂) end @@ -343,14 +343,14 @@ section -- 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 (take d, assume H : a * c = a * b * d, have H1 : b * d = c, from mul.cancel_left Ha (mul.assoc a b d ▸ H⁻¹), 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 (take d, assume H : c * a = b * a * d, diff --git a/library/data/int/basic.lean b/library/data/int/basic.lean index 01868e2b8..76cafb916 100644 --- a/library/data/int/basic.lean +++ b/library/data/int/basic.lean @@ -710,35 +710,35 @@ section port_algebra theorem ne_zero_of_mul_ne_zero_left : ∀{a b : ℤ}, a * b ≠ 0 → b ≠ 0 := @algebra.ne_zero_of_mul_ne_zero_left _ _ definition dvd (a b : ℤ) : Prop := algebra.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_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 := + notation a ∣ b := dvd a b + 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 exists_eq_mul_right_of_dvd : ∀{a b : ℤ} (H : a ∣ b), ∃c, b = a * c := @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 _ _ - 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 _ _ - 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 _ _ - 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 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 one_dvd : ∀a : ℤ, (1 | a) := algebra.one_dvd - 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_of_dvd_left : ∀{a b : ℤ} (H : (a | b)) (c : ℤ), (a | b * c) := + 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 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 one_dvd : ∀a : ℤ, 1 ∣ a := algebra.one_dvd + 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_of_dvd_left : ∀{a b : ℤ} (H : a ∣ b) (c : ℤ), a ∣ b * c := @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 _ _ - 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 _ _ - 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 _ _ - 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 _ _ - 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 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 @@ -757,9 +757,9 @@ section port_algebra algebra.mul_self_sub_mul_self_eq theorem mul_self_sub_one_eq : ∀a : ℤ, a * a - 1 = (a + 1) * (a - 1) := algebra.mul_self_sub_one_eq - 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 dvd_sub : ∀a b c : ℤ, (a | b) → (a | c) → (a | b - c) := algebra.dvd_sub + 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 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.cancel_right : ∀{a b c : ℤ}, a ≠ 0 → b * a = c * a → b = c := @algebra.mul.cancel_right _ _ @@ -769,9 +769,9 @@ section port_algebra algebra.mul_self_eq_mul_self_iff theorem mul_self_eq_one_iff : ∀a : ℤ, a * a = 1 ↔ a = 1 ∨ a = -1 := 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 _ _ - 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 _ _ end port_algebra diff --git a/library/data/int/order.lean b/library/data/int/order.lean index d4686ba2d..a81118c08 100644 --- a/library/data/int/order.lean +++ b/library/data/int/order.lean @@ -535,8 +535,8 @@ section port_algebra 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 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 dvd_abs_iff : ∀a b : ℤ, (a | abs b) ↔ (a | b) := algebra.dvd_abs_iff + 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 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 end port_algebra diff --git a/library/data/nat/basic.lean b/library/data/nat/basic.lean index 1b0d87039..784ffce01 100644 --- a/library/data/nat/basic.lean +++ b/library/data/nat/basic.lean @@ -315,36 +315,36 @@ section port_algebra 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 - 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_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 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 exists_eq_mul_right_of_dvd : ∀{a b : ℕ} (H : a ∣ b), ∃c, b = a * c := @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 _ _ - 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 _ _ - 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 _ _ - 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 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 one_dvd : ∀a : ℕ, (1 | a) := algebra.one_dvd - 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_of_dvd_left : ∀{a b : ℕ} (H : (a | b)) (c : ℕ), (a | b * c) := + theorem dvd.refl : ∀a : ℕ, a ∣ a := algebra.dvd.refl + theorem dvd.trans : ∀{a b c : ℕ}, a ∣ b → b ∣ c → a ∣ c := @algebra.dvd.trans _ _ + 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 one_dvd : ∀a : ℕ, 1 ∣ a := algebra.one_dvd + 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_of_dvd_left : ∀{a b : ℕ} (H : a ∣ b) (c : ℕ), a ∣ b * c := @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 _ _ - 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 _ _ - 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 _ _ - 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 _ _ - 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 nat diff --git a/library/data/nat/div.lean b/library/data/nat/div.lean index 22ef5d148..d85d96cdf 100644 --- a/library/data/nat/div.lean +++ b/library/data/nat/div.lean @@ -332,28 +332,28 @@ calc /- 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) -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 (take z, assume H1 : n = m * z, 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 definition dvd.decidable_rel [instance] : decidable_rel dvd := 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) -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 -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 m * k = m * (n div m) : H2 ... = n : mul_div_cancel' H1) @@ -363,7 +363,7 @@ calc n = n * k div k : mul_div_cancel _ H1 ... = 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₁ = m * c₂), from H₂, 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, 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) -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 (assume H3 : n1 ≥ n2, 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), 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 - (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, assume Hpos : n > 0, - assume H1 : (m | n), - assume H2 : (n | m), + assume H1 : m ∣ n, + assume H2 : n ∣ m, 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, 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, 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) (assume H1 : k = 0, 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) : 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 (take l, assume H1 : k * n = k * m * l, have H2 : n = m * l, from eq_of_mul_eq_mul_left kpos (H1 ⬝ !mul.assoc), 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) -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 H4 : n = n div k * k, from (div_mul_cancel (dvd.trans H1 H2))⁻¹, or.elim (eq_zero_or_pos k) @@ -518,36 +518,36 @@ have aux : Q (m, n), from H1 m (succ n₁) hlt₁ hp))), 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 (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, assume npos : 0 < 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 + 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 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⁻¹, - 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 - (take m, assume (h₁ : (k | m)) (h₂ : (k | 0)), - show (k | gcd m 0), from !gcd_zero_right⁻¹ ▸ h₁) + (take m, assume (h₁ : k ∣ m) (h₂ : k ∣ 0), + show k ∣ gcd m 0, from !gcd_zero_right⁻¹ ▸ h₁) (take m n, assume npos : n > 0, - assume IH : (k | n) → (k | m mod n) → (k | gcd n (m mod n)), - assume H1 : (k | m), - assume H2 : (k | n), - 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)), + assume IH : k ∣ n → k ∣ m mod n → k ∣ gcd n (m mod n), + assume H1 : k ∣ m, + assume H2 : k ∣ n, + 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 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 := 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 := 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) (assume H3 : k = 0, 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_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 -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 -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) -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 /- lcm -/ @@ -669,17 +669,17 @@ calc ... = m * m div m : gcd_self ... = 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, 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 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) -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) (assume kzero : k = 0, !kzero⁻¹ ▸ !dvd_zero) (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 := !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 calc gcd (m * k) (m * n) = m * gcd k n : gcd_mul_left ... = m * 1 : H1 ... = 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 -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) 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)) 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) 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 := coprime_of_coprime_mul_left_right (!mul.comm ▸ H) -theorem exists_eq_prod_and_dvd_and_dvd {m n k} (H : (k | m * n)) : - ∃ m' n', k = m' * n' ∧ (m' | m) ∧ (n' | 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 := or.elim (eq_zero_or_pos (gcd k m)) (assume H1 : gcd k m = 0, 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 H4 : k = 0 * n, from H2 ⬝ !zero_mul⁻¹, - have H5 : (0 | m), from H3⁻¹ ▸ !dvd.refl, - have H6 : (n | n), from !dvd.refl, + have H5 : 0 ∣ m, from H3⁻¹ ▸ !dvd.refl, + have H6 : n ∣ n, from !dvd.refl, exists.intro _ (exists.intro _ (and.intro H4 (and.intro H5 H6)))) (assume H1 : gcd k m > 0, - 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 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 H4 : (m * n) div gcd k m = (m div gcd k m) * n, from calc 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 ... = 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 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)⁻¹, exists.intro _ (exists.intro _ (and.intro H8 (and.intro !gcd_dvd_right H7)))) diff --git a/library/data/nat/order.lean b/library/data/nat/order.lean index 3e115ab9c..598541cb9 100644 --- a/library/data/nat/order.lean +++ b/library/data/nat/order.lean @@ -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 := 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 (assume H3 : m = 0, 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 := 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 (take m, assume H1 : 1 = n * m, diff --git a/library/init/reserved_notation.lean b/library/init/reserved_notation.lean index 3c17d9d48..d72424a81 100644 --- a/library/init/reserved_notation.lean +++ b/library/init/reserved_notation.lean @@ -89,7 +89,7 @@ reserve infixl `∪`:65 /- other symbols -/ -reserve notation `(` a `|` b `)` +reserve infix `∣`:50 reserve infixl `++`:65 reserve infixr `::`:65 diff --git a/library/init/tactic.lean b/library/init/tactic.lean index bf5195651..6262d15b1 100644 --- a/library/init/tactic.lean +++ b/library/init/tactic.lean @@ -92,13 +92,13 @@ opaque definition change (e : expr) : tactic := builtin opaque definition assert_hypothesis (id : expr) (e : expr) : tactic := builtin 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 focus (t : tactic) : tactic := focus_at t 0 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 := nat.rec id (λn t', (t;t')) (nat.of_num n) diff --git a/src/frontends/lean/parser.cpp b/src/frontends/lean/parser.cpp index 2cc98987f..d3ffa1527 100644 --- a/src/frontends/lean/parser.cpp +++ b/src/frontends/lean/parser.cpp @@ -1403,7 +1403,11 @@ expr parser::parse_tactic_nud() { } else { 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(); expr r = parse_tactic(); while (curr_is_token(get_bar_tk())) { @@ -1412,15 +1416,6 @@ expr parser::parse_tactic_nud() { expr n = parse_tactic(); 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"); return r; } else { diff --git a/tests/lean/run/assert_tac2.lean b/tests/lean/run/assert_tac2.lean index f643eadc9..d2e24e282 100644 --- a/tests/lean/run/assert_tac2.lean +++ b/tests/lean/run/assert_tac2.lean @@ -1,7 +1,7 @@ import data.nat 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 | @or.inl _ _ kzero := begin diff --git a/tests/lean/run/class7.lean b/tests/lean/run/class7.lean index ae001064f..7453b1dd9 100644 --- a/tests/lean/run/class7.lean +++ b/tests/lean/run/class7.lean @@ -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 T1 {A : Type} (a : A) : inh A := -by repeat [apply @inh.intro | eassumption] +by repeat (apply @inh.intro | eassumption) theorem T2 : inh Prop diff --git a/tests/lean/run/is_true.lean b/tests/lean/run/is_true.lean index 444a2c0c5..537c8a122 100644 --- a/tests/lean/run/is_true.lean +++ b/tests/lean/run/is_true.lean @@ -4,5 +4,5 @@ open nat example : is_true (2 = 2) := trivial example : is_false (3 = 2) := trivial example : is_true (2 < 3) := trivial -example : is_true (3 | 9) := trivial -example : is_false (3 | 7) := trivial +example : is_true (3 ∣ 9) := trivial +example : is_false (3 ∣ 7) := trivial diff --git a/tests/lean/run/match_tac4.lean b/tests/lean/run/match_tac4.lean index 6ec665061..62944d07e 100644 --- a/tests/lean/run/match_tac4.lean +++ b/tests/lean/run/match_tac4.lean @@ -16,7 +16,7 @@ begin match H with | ⟪ H₁, H₂, H₃, H₄ ⟫ := begin - repeat [apply and.intro | assumption] + repeat (apply and.intro | assumption) 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 := begin 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 print definition tst2 diff --git a/tests/lean/run/tactic14.lean b/tests/lean/run/tactic14.lean index dd52822cf..1567b7e84 100644 --- a/tests/lean/run/tactic14.lean +++ b/tests/lean/run/tactic14.lean @@ -2,7 +2,7 @@ import logic open 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 diff --git a/tests/lean/run/tactic22.lean b/tests/lean/run/tactic22.lean index ae413ac8a..9400e0378 100644 --- a/tests/lean/run/tactic22.lean +++ b/tests/lean/run/tactic22.lean @@ -2,4 +2,4 @@ import logic open tactic 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)) diff --git a/tests/lean/run/tactic24.lean b/tests/lean/run/tactic24.lean index 932651064..e64aca3d5 100644 --- a/tests/lean/run/tactic24.lean +++ b/tests/lean/run/tactic24.lean @@ -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 := _ -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 | - assumption]) + assumption)) tactic_hint my_tac3 diff --git a/tests/lean/run/tactic25.lean b/tests/lean/run/tactic25.lean index 2d06a5ac7..e6fc9d682 100644 --- a/tests/lean/run/tactic25.lean +++ b/tests/lean/run/tactic25.lean @@ -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 -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 | - assumption]) + assumption)) tactic_hint my_tac3 theorem T3 {a b c : Prop} (Hb : b) : a ∨ b ∨ c diff --git a/tests/lean/run/tactic26.lean b/tests/lean/run/tactic26.lean index c6aeca601..7f86e364d 100644 --- a/tests/lean/run/tactic26.lean +++ b/tests/lean/run/tactic26.lean @@ -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) := 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 @num.is_inhabited - ]) + )) tactic_hint my_tac diff --git a/tests/lean/run/tactic27.lean b/tests/lean/run/tactic27.lean index 560471864..6eff6ea7f 100644 --- a/tests/lean/run/tactic27.lean +++ b/tests/lean/run/tactic27.lean @@ -1,9 +1,7 @@ import logic open tactic -definition my_tac := repeat ([ apply @and.intro - | apply @eq.refl - ]) +definition my_tac := repeat (apply @and.intro | apply @eq.refl) tactic_hint my_tac theorem T1 {A : Type} {B : Type} (a : A) (b c : B) : a = a ∧ b = b ∧ c = c diff --git a/tests/lean/run/tactic3.lean b/tests/lean/run/tactic3.lean index 6ecc69c03..55e0e576f 100644 --- a/tests/lean/run/tactic3.lean +++ b/tests/lean/run/tactic3.lean @@ -2,8 +2,8 @@ import logic open tactic 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 "third"; assumption] + trace "third"; assumption) check tst diff --git a/tests/lean/run/tactic7.lean b/tests/lean/run/tactic7.lean index ae8152331..cb7aed129 100644 --- a/tests/lean/run/tactic7.lean +++ b/tests/lean/run/tactic7.lean @@ -6,6 +6,6 @@ theorem tst {A B : Prop} (H1 : A) (H2 : B) : A ∧ B ∧ A check tst 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