diff --git a/hott/algebra/category/basic.hlean b/hott/algebra/category/basic.hlean index 91ba7f3cc..85ffbbe85 100644 --- a/hott/algebra/category/basic.hlean +++ b/hott/algebra/category/basic.hlean @@ -43,7 +43,7 @@ namespace category definition is_trunc_1_ob : is_trunc 1 ob := begin - apply is_trunc_succ_intro, intros (a, b), + apply is_trunc_succ_intro, intros [a, b], fapply is_trunc_is_equiv_closed, exact (@eq_of_iso _ _ a b), apply is_equiv_inv, diff --git a/hott/algebra/category/constructions.hlean b/hott/algebra/category/constructions.hlean index 7cbf70231..d715644c9 100644 --- a/hott/algebra/category/constructions.hlean +++ b/hott/algebra/category/constructions.hlean @@ -95,7 +95,7 @@ namespace category begin fapply functor_eq, {exact (eq_of_iso_functor_ob η)}, - {intros (c, c', f), --unfold eq_of_iso_functor_ob, --TODO: report: this fails + {intros [c, c', f], --unfold eq_of_iso_functor_ob, --TODO: report: this fails apply concat, {apply (ap (λx, to_hom x ∘ to_fun_hom F f ∘ _)), apply (retr iso_of_eq)}, apply concat, @@ -108,8 +108,8 @@ namespace category apply iso.eq_mk, apply nat_trans_eq_mk, intro c, - rewrite natural_map_hom_of_eq, esimp {eq_of_iso_functor}, - rewrite ap010_functor_eq, esimp {hom_of_eq,eq_of_iso_functor_ob}, + rewrite natural_map_hom_of_eq, esimp [eq_of_iso_functor], + rewrite ap010_functor_eq, esimp [hom_of_eq,eq_of_iso_functor_ob], rewrite (retr iso_of_eq), end @@ -117,9 +117,9 @@ namespace category begin apply functor_eq2, intro c, - esimp {eq_of_iso_functor}, + esimp [eq_of_iso_functor], rewrite ap010_functor_eq, - esimp {eq_of_iso_functor_ob}, + esimp [eq_of_iso_functor_ob], rewrite componentwise_iso_iso_of_eq, rewrite (sect iso_of_eq) end diff --git a/hott/algebra/groupoid.hlean b/hott/algebra/groupoid.hlean index 591c9b7b4..3cea3268a 100644 --- a/hott/algebra/groupoid.hlean +++ b/hott/algebra/groupoid.hlean @@ -42,9 +42,9 @@ namespace category [G : groupoid ob] : group (hom (center ob) (center ob)) := begin fapply group.mk, - intros (f, g), apply (comp f g), + intros [f, g], apply (comp f g), apply homH, - intros (f, g, h), apply (assoc f g h)⁻¹, + intros [f, g, h], apply (assoc f g h)⁻¹, apply (ID (center ob)), intro f, apply id_left, intro f, apply id_right, @@ -55,9 +55,9 @@ namespace category definition group_of_groupoid_unit [G : groupoid unit] : group (hom ⋆ ⋆) := begin fapply group.mk, - intros (f, g), apply (comp f g), + intros [f, g], apply (comp f g), apply homH, - intros (f, g, h), apply (assoc f g h)⁻¹, + intros [f, g, h], apply (assoc f g h)⁻¹, apply (ID ⋆), intro f, apply id_left, intro f, apply id_right, @@ -71,7 +71,7 @@ namespace category fapply groupoid.mk, intros, exact A, intros, apply (@group.carrier_hset A G), - intros (a, b, c, g, h), exact (@group.mul A G g h), + intros [a, b, c, g, h], exact (@group.mul A G g h), intro a, exact (@group.one A G), intros, exact (@group.mul_assoc A G h g f)⁻¹, intros, exact (@group.one_mul A G f), @@ -85,9 +85,9 @@ namespace category group (hom a a) := begin fapply group.mk, - intros (f, g), apply (comp f g), + intros [f, g], apply (comp f g), apply homH, - intros (f, g, h), apply (assoc f g h)⁻¹, + intros [f, g, h], apply (assoc f g h)⁻¹, apply (ID a), intro f, apply id_left, intro f, apply id_right, diff --git a/hott/algebra/precategory/constructions.hlean b/hott/algebra/precategory/constructions.hlean index 3c877c937..67f76c39a 100644 --- a/hott/algebra/precategory/constructions.hlean +++ b/hott/algebra/precategory/constructions.hlean @@ -41,7 +41,7 @@ namespace category -- symmetric associativity proof. definition opposite_opposite' {ob : Type} (C : precategory ob) : opposite (opposite C) = C := begin - apply (precategory.rec_on C), intros (hom', homH', comp', ID', assoc', id_left', id_right'), + apply (precategory.rec_on C), intros [hom', homH', comp', ID', assoc', id_left', id_right'], apply (ap (λassoc'', precategory.mk hom' @homH' comp' ID' assoc'' id_left' id_right')), repeat (apply eq_of_homotopy ; intros ), apply ap, diff --git a/hott/algebra/precategory/functor.hlean b/hott/algebra/precategory/functor.hlean index 1aa02d6cb..147698ccb 100644 --- a/hott/algebra/precategory/functor.hlean +++ b/hott/algebra/precategory/functor.hlean @@ -136,14 +136,14 @@ namespace functor exact (S.1), exact (S.2.1), exact (pr₁ S.2.2), exact (pr₂ S.2.2)}, {intro F, - cases F with (d1, d2, d3, d4), + cases F with [d1, d2, d3, d4], exact ⟨d1, d2, (d3, @d4)⟩}, {intro F, cases F, apply idp}, {intro S, - cases S with (d1, S2), - cases S2 with (d2, P1), + cases S with [d1, S2], + cases S2 with [d2, P1], cases P1, apply idp}, end @@ -192,8 +192,8 @@ namespace functor definition functor_eq2 {F₁ F₂ : C ⇒ D} (p q : F₁ = F₂) (r : ap010 to_fun_ob p ∼ ap010 to_fun_ob q) : p = q := begin - cases F₁ with (ob₁, hom₁, id₁, comp₁), - cases F₂ with (ob₂, hom₂, id₂, comp₂), + cases F₁ with [ob₁, hom₁, id₁, comp₁], + cases F₂ with [ob₂, hom₂, id₂, comp₂], rewrite [-functor_eq_eta' p, -functor_eq_eta' q], apply functor_eq2', apply ap_eq_ap_of_homotopy, @@ -213,7 +213,7 @@ namespace functor (q : (λ(a b : C) (f : hom a b), hom_of_eq (p b) ∘ F₁ f ∘ inv_of_eq (p a)) ∼3 to_fun_hom F₂) (c : C) : ap010 to_fun_ob (functor_eq p q) c = p c := begin - cases F₂, revert q, apply (homotopy.rec_on p), clear p, esimp, intros (p, q), + cases F₂, revert q, apply (homotopy.rec_on p), clear p, esimp, intros [p, q], apply sorry, --apply (homotopy3.rec_on q), clear q, intro q, --cases p, --TODO: report: this fails diff --git a/hott/algebra/precategory/iso.hlean b/hott/algebra/precategory/iso.hlean index c12309f05..9c59835e9 100644 --- a/hott/algebra/precategory/iso.hlean +++ b/hott/algebra/precategory/iso.hlean @@ -129,8 +129,8 @@ namespace iso definition is_hprop_is_iso [instance] (f : hom a b) : is_hprop (is_iso f) := begin - apply is_hprop.mk, intros (H, H'), - cases H with (g, li, ri), cases H' with (g', li', ri'), + apply is_hprop.mk, intros [H, H'], + cases H with [g, li, ri], cases H' with [g', li', ri'], fapply (apD0111 (@is_iso.mk ob C a b f)), apply left_inverse_eq_right_inverse, apply li, @@ -178,7 +178,7 @@ namespace iso fapply (equiv.mk), {intro S, apply iso.mk, apply (S.2)}, {fapply adjointify, - {intro p, cases p with (f, H), exact (sigma.mk f H)}, + {intro p, cases p with [f, H], exact (sigma.mk f H)}, {intro p, cases p, apply idp}, {intro S, cases S, apply idp}}, end diff --git a/hott/algebra/precategory/nat_trans.hlean b/hott/algebra/precategory/nat_trans.hlean index 2b57221b6..09237859c 100644 --- a/hott/algebra/precategory/nat_trans.hlean +++ b/hott/algebra/precategory/nat_trans.hlean @@ -68,7 +68,7 @@ namespace nat_trans intro H, fapply sigma.mk, intro a, exact (H a), - intros (a, b, f), exact (naturality H f), + intros [a, b, f], exact (naturality H f), intro η, apply nat_trans_eq_mk, intro a, apply idp, intro S, fapply sigma_eq, diff --git a/hott/algebra/precategory/yoneda.hlean b/hott/algebra/precategory/yoneda.hlean index 32de74b00..071c39631 100644 --- a/hott/algebra/precategory/yoneda.hlean +++ b/hott/algebra/precategory/yoneda.hlean @@ -31,7 +31,7 @@ namespace yoneda intro x, apply eq_of_homotopy, intro h, exact (!id_left ⬝ !id_right) end begin - intros (x, y, z, g, f), apply eq_of_homotopy, intro h, + intros [x, y, z, g, f], apply eq_of_homotopy, intro h, exact (representable_functor_assoc g.2 f.2 h f.1 g.1), end end yoneda @@ -128,8 +128,8 @@ namespace functor theorem functor_uncurry_functor_curry : functor_uncurry (functor_curry F) = F := functor_eq (λp, ap (to_fun_ob F) !prod.eta) begin - intros (cd, cd', fg), - cases cd with (c,d), cases cd' with (c',d'), cases fg with (f,g), + intros [cd, cd', fg], + cases cd with [c,d], cases cd' with [c',d'], cases fg with [f, g], apply concat, apply id_leftright, show (functor_uncurry (functor_curry F)) (f, g) = F (f,g), from calc @@ -144,7 +144,7 @@ namespace functor begin fapply functor_eq, {intro d, apply idp}, - {intros (d, d', g), + {intros [d, d', g], apply concat, apply id_leftright, show to_fun_hom (functor_curry (functor_uncurry G) c) g = to_fun_hom (G c) g, from calc @@ -158,7 +158,7 @@ namespace functor theorem functor_curry_functor_uncurry : functor_curry (functor_uncurry G) = G := begin fapply functor_eq, exact (functor_curry_functor_uncurry_ob G), - intros (c, c', f), + intros [c, c', f], fapply nat_trans_eq_mk, intro d, apply concat, @@ -191,7 +191,7 @@ namespace functor : functor_prod_flip D C ∘f (functor_prod_flip C D) = functor.id := begin fapply functor_eq, {intro p, apply prod.eta}, - intros (p, p', h), cases p with (c, d), cases p' with (c', d'), + intros [p, p', h], cases p with [c, d], cases p' with [c', d'], apply id_leftright, end end functor diff --git a/hott/arity.hlean b/hott/arity.hlean index 58ed580f1..cac814dbe 100644 --- a/hott/arity.hlean +++ b/hott/arity.hlean @@ -132,7 +132,7 @@ namespace funext adjointify _ eq_of_homotopy2 begin - intro H, esimp {apD100,eq_of_homotopy2, function.compose}, + intro H, esimp [apD100, eq_of_homotopy2, function.compose], apply eq_of_homotopy, intro a, apply concat, apply (ap (λx, @apD10 _ (λb : B a, _) _ _ (x a))), apply (retr apD10), --TODO: remove implicit argument after #469 is closed diff --git a/hott/init/nat.hlean b/hott/init/nat.hlean index 61356f992..8c2fbd57d 100644 --- a/hott/init/nat.hlean +++ b/hott/init/nat.hlean @@ -152,7 +152,7 @@ namespace nat definition le.rec_on {a : nat} {P : nat → Type} {b : nat} (H : a ≤ b) (H₁ : P a) (H₂ : ∀ b, a < b → P b) : P b := begin - cases H with (b', hlt), + cases H with [b', hlt], apply H₁, apply (H₂ b hlt) end @@ -201,21 +201,21 @@ namespace nat definition le.trans {a b c : nat} (h₁ : a ≤ b) (h₂ : b ≤ c) : a ≤ c := begin - cases h₁ with (b', hlt), + cases h₁ with [b', hlt], apply h₂, apply (lt.trans hlt h₂) end definition lt.of_le_of_lt {a b c : nat} (h₁ : a ≤ b) (h₂ : b < c) : a < c := begin - cases h₁ with (b', hlt), + cases h₁ with [b', hlt], apply h₂, apply (lt.trans hlt h₂) end definition lt.of_lt_of_le {a b c : nat} (h₁ : a < b) (h₂ : b ≤ c) : a < c := begin - cases h₁ with (b', hlt), + cases h₁ with [b', hlt], apply (lt.of_succ_lt_succ h₂), apply (lt.trans hlt (lt.of_succ_lt_succ h₂)) end diff --git a/hott/types/equiv.hlean b/hott/types/equiv.hlean index 1dc62f774..8b049fd57 100644 --- a/hott/types/equiv.hlean +++ b/hott/types/equiv.hlean @@ -62,9 +62,9 @@ namespace is_equiv equiv.MK (λH, ⟨inv f, retr f, sect f, adj f⟩) (λp, is_equiv.mk p.1 p.2.1 p.2.2.1 p.2.2.2) (λp, begin - cases p with (p1, p2), - cases p2 with (p21, p22), - cases p22 with (p221, p222), + cases p with [p1, p2], + cases p2 with [p21, p22], + cases p22 with [p221, p222], apply idp end) (λH, is_equiv.rec_on H (λH1 H2 H3 H4, idp)) diff --git a/hott/types/pi.hlean b/hott/types/pi.hlean index 75701a9a3..3d69832e0 100644 --- a/hott/types/pi.hlean +++ b/hott/types/pi.hlean @@ -157,15 +157,15 @@ namespace pi definition is_trunc_pi [instance] (B : A → Type) (n : trunc_index) [H : ∀a, is_trunc n (B a)] : is_trunc n (Πa, B a) := begin - reverts (B, H), + reverts [B, H], apply (trunc_index.rec_on n), - {intros (B, H), + {intros [B, H], fapply is_contr.mk, intro a, apply center, intro f, apply eq_of_homotopy, intro x, apply (contr (f x))}, - {intros (n, IH, B, H), - fapply is_trunc_succ_intro, intros (f, g), + {intros [n, IH, B, H], + fapply is_trunc_succ_intro, intros [f, g], fapply is_trunc_equiv_closed, apply equiv.symm, apply eq_equiv_homotopy, apply IH, diff --git a/hott/types/prod.hlean b/hott/types/prod.hlean index 6fc6197a1..ef1e5a200 100644 --- a/hott/types/prod.hlean +++ b/hott/types/prod.hlean @@ -25,8 +25,8 @@ namespace prod definition prod_eq (H₁ : pr₁ u = pr₁ v) (H₂ : pr₂ u = pr₂ v) : u = v := begin - cases u with (a₁, b₁), - cases v with (a₂, b₂), + cases u with [a₁, b₁], + cases v with [a₂, b₂], apply (transport _ (eta (a₁, b₁))), apply (transport _ (eta (a₂, b₂))), apply (pair_eq H₁ H₂), diff --git a/hott/types/sigma.hlean b/hott/types/sigma.hlean index 6bae14389..8bdaa1b11 100644 --- a/hott/types/sigma.hlean +++ b/hott/types/sigma.hlean @@ -146,9 +146,9 @@ namespace sigma definition sigma_eq2 {p q : u = v} (r : p..1 = q..1) (s : r ▹ p..2 = q..2) : p = q := begin - reverts (q, r, s), - cases p, cases u with (u1, u2), - intros (q, r, s), + reverts [q, r, s], + cases p, cases u with [u1, u2], + intros [q, r, s], apply concat, rotate 1, apply sigma_eq_eta, apply (sigma_eq_eq_sigma_eq r s) @@ -179,7 +179,7 @@ namespace sigma definition tr_eq2_nondep {C : A → Type} {D : Π a:A, B a → C a → Type} (p : a = a') (bcd : Σ(b : B a) (c : C a), D a b c) : p ▹ bcd = ⟨p ▹ bcd.1, p ▹ bcd.2.1, p ▹D2 bcd.2.2⟩ := begin - cases p, cases bcd with (b, cd), + cases p, cases bcd with [b, cd], cases cd, apply idp end @@ -198,7 +198,7 @@ namespace sigma ((g (f⁻¹ a'))⁻¹ (transport B' (retr f a')⁻¹ b')))) begin intro u', - cases u' with (a', b'), + cases u' with [a', b'], apply (sigma_eq (retr f a')), -- rewrite retr, -- end @@ -209,7 +209,7 @@ namespace sigma end begin intro u, - cases u with (a, b), + cases u with [a, b], apply (sigma_eq (sect f a)), show transport B (sect f a) ((g (f⁻¹ (f a)))⁻¹ (transport B' (retr f (f a))⁻¹ (g a b))) = b, from calc @@ -280,8 +280,8 @@ namespace sigma equiv.mk _ (adjointify (λav, ⟨⟨av.1, av.2.1⟩, av.2.2⟩) (λuc, ⟨uc.1.1, uc.1.2, !eta⁻¹ ▹ uc.2⟩) - begin intro uc; cases uc with (u, c); cases u; apply idp end - begin intro av; cases av with (a, v); cases v; apply idp end) + begin intro uc; cases uc with [u, c]; cases u; apply idp end + begin intro av; cases av with [a, v]; cases v; apply idp end) open prod definition assoc_equiv_prod (C : (A × A') → Type) : (Σa a', C (a,a')) ≃ (Σu, C u) := @@ -369,14 +369,14 @@ namespace sigma definition is_trunc_sigma (B : A → Type) (n : trunc_index) [HA : is_trunc n A] [HB : Πa, is_trunc n (B a)] : is_trunc n (Σa, B a) := begin - reverts (A, B, HA, HB), + reverts [A, B, HA, HB], apply (trunc_index.rec_on n), - intros (A, B, HA, HB), + intros [A, B, HA, HB], fapply is_trunc.is_trunc_equiv_closed, apply equiv.symm, apply sigma_equiv_of_is_contr_pr1, - intros (n, IH, A, B, HA, HB), - fapply is_trunc.is_trunc_succ_intro, intros (u, v), + intros [n, IH, A, B, HA, HB], + fapply is_trunc.is_trunc_succ_intro, intros [u, v], fapply is_trunc.is_trunc_equiv_closed, apply equiv_sigma_eq, apply IH, diff --git a/hott/types/trunc.hlean b/hott/types/trunc.hlean index a7c0e05c8..5a741e827 100644 --- a/hott/types/trunc.hlean +++ b/hott/types/trunc.hlean @@ -21,7 +21,7 @@ namespace is_trunc {fapply is_equiv.adjointify, {intro H, apply sigma.mk, exact (@contr A H)}, {intro H, apply (is_trunc.rec_on H), intro Hint, - apply (contr_internal.rec_on Hint), intros (H1, H2), + apply (contr_internal.rec_on Hint), intros [H1, H2], apply idp}, {intro S, cases S, apply idp}} end @@ -31,10 +31,10 @@ namespace is_trunc begin fapply equiv.MK, {intro H, apply is_trunc_succ_intro}, - {intros (H, x, y), apply is_trunc_eq}, + {intros [H, x, y], apply is_trunc_eq}, {intro H, apply (is_trunc.rec_on H), intro Hint, apply idp}, {intro P, apply eq_of_homotopy, intro a, apply eq_of_homotopy, intro b, - esimp {function.id,compose,is_trunc_succ_intro,is_trunc_eq}, + esimp [function.id,compose,is_trunc_succ_intro,is_trunc_eq], generalize (P a b), intro H, apply (is_trunc.rec_on H), intro H', apply idp}, end @@ -49,14 +49,14 @@ namespace is_trunc fapply sigma_eq, apply x.2, apply (@is_hprop.elim), apply is_trunc_pi, intro a, - apply is_hprop.mk, intros (w, z), + apply is_hprop.mk, intros [w, z], have H : is_hset A, begin apply is_trunc_succ, apply is_trunc_succ, apply is_contr.mk, exact y.2 end, fapply (@is_hset.elim A _ _ _ w z)}, - {intros (n', IH, A), + {intros [n', IH, A], apply is_trunc_is_equiv_closed, apply equiv.to_is_equiv, apply is_trunc.pi_char}, @@ -124,8 +124,8 @@ namespace is_trunc fapply equiv.MK, /--/ intro A, exact (⟨carrier A, struct A⟩), /--/ intro S, exact (trunctype.mk S.1 S.2), - /--/ intro S, apply (sigma.rec_on S), intros (S1, S2), apply idp, - intro A, apply (trunctype.rec_on A), intros (A1, A2), apply idp, + /--/ intro S, apply (sigma.rec_on S), intros [S1, S2], apply idp, + intro A, apply (trunctype.rec_on A), intros [A1, A2], apply idp, end -- set_option pp.notation false diff --git a/library/algebra/group.lean b/library/algebra/group.lean index eafe3b1b8..c8246d307 100644 --- a/library/algebra/group.lean +++ b/library/algebra/group.lean @@ -151,20 +151,20 @@ section group theorem mul.left_inv (a : A) : a⁻¹ * a = 1 := !group.mul_left_inv theorem inv_mul_cancel_left (a b : A) : a⁻¹ * (a * b) = b := - by rewrite ⟨-mul.assoc, mul.left_inv, one_mul⟩ + by rewrite [-mul.assoc, mul.left_inv, one_mul] theorem inv_mul_cancel_right (a b : A) : a * b⁻¹ * b = a := - by rewrite ⟨mul.assoc, mul.left_inv, mul_one⟩ + by rewrite [mul.assoc, mul.left_inv, mul_one] theorem inv_eq_of_mul_eq_one {a b : A} (H : a * b = 1) : a⁻¹ = b := - by rewrite ⟨-mul_one a⁻¹, -H, inv_mul_cancel_left⟩ + by rewrite [-mul_one a⁻¹, -H, inv_mul_cancel_left] theorem inv_one : 1⁻¹ = 1 := inv_eq_of_mul_eq_one (one_mul 1) theorem inv_inv (a : A) : (a⁻¹)⁻¹ = a := inv_eq_of_mul_eq_one (mul.left_inv a) theorem inv.inj {a b : A} (H : a⁻¹ = b⁻¹) : a = b := - by rewrite ⟨-inv_inv, H, inv_inv⟩ + by rewrite [-inv_inv, H, inv_inv] theorem inv_eq_inv_iff_eq (a b : A) : a⁻¹ = b⁻¹ ↔ a = b := iff.intro (assume H, inv.inj H) (assume H, congr_arg _ H) @@ -239,10 +239,10 @@ section group iff.intro eq_mul_inv_of_mul_eq mul_eq_of_eq_mul_inv theorem mul_left_cancel {a b c : A} (H : a * b = a * c) : b = c := - by rewrite ⟨-inv_mul_cancel_left a b, H, inv_mul_cancel_left⟩ + by rewrite [-inv_mul_cancel_left a b, H, inv_mul_cancel_left] theorem mul_right_cancel {a b c : A} (H : a * b = c * b) : a = c := - by rewrite ⟨-mul_inv_cancel_right a b, H, mul_inv_cancel_right⟩ + by rewrite [-mul_inv_cancel_right a b, H, mul_inv_cancel_right] definition group.to_left_cancel_semigroup [instance] [coercion] [reducible] : left_cancel_semigroup A := ⦃ left_cancel_semigroup, s, @@ -269,13 +269,13 @@ section add_group theorem add.left_inv (a : A) : -a + a = 0 := !add_group.add_left_inv theorem neg_add_cancel_left (a b : A) : -a + (a + b) = b := - by rewrite ⟨-add.assoc, add.left_inv, zero_add⟩ + by rewrite [-add.assoc, add.left_inv, zero_add] theorem neg_add_cancel_right (a b : A) : a + -b + b = a := - by rewrite ⟨add.assoc, add.left_inv, add_zero⟩ + by rewrite [add.assoc, add.left_inv, add_zero] theorem neg_eq_of_add_eq_zero {a b : A} (H : a + b = 0) : -a = b := - by rewrite ⟨-add_zero, -H, neg_add_cancel_left⟩ + by rewrite [-add_zero, -H, neg_add_cancel_left] theorem neg_zero : -0 = 0 := neg_eq_of_add_eq_zero (zero_add 0) @@ -304,15 +304,15 @@ section add_group ... = 0 : add.left_inv theorem add_neg_cancel_left (a b : A) : a + (-a + b) = b := - by rewrite ⟨-add.assoc, add.right_inv, zero_add⟩ + by rewrite [-add.assoc, add.right_inv, zero_add] theorem add_neg_cancel_right (a b : A) : a + b + -b = a := - by rewrite ⟨add.assoc, add.right_inv, add_zero⟩ + by rewrite [add.assoc, add.right_inv, add_zero] theorem neg_add_rev (a b : A) : -(a + b) = -b + -a := neg_eq_of_add_eq_zero begin - rewrite ⟨add.assoc, add_neg_cancel_left, add.right_inv⟩ + rewrite [add.assoc, add_neg_cancel_left, add.right_inv] end -- TODO: delete these in favor of sub rules? @@ -454,7 +454,7 @@ section add_comm_group by rewrite [▸ a + -b + -c = _, add.assoc, -neg_add] theorem add_sub_add_left_eq_sub (a b c : A) : (c + a) - (c + b) = a - b := - by rewrite ⟨sub_add_eq_sub_sub, (add.comm c a), add_sub_cancel⟩ + by rewrite [sub_add_eq_sub_sub, (add.comm c a), add_sub_cancel] theorem eq_sub_of_add_eq' {a b c : A} (H : c + a = b) : a = b - c := !eq_sub_of_add_eq (!add.comm ▸ H) diff --git a/library/algebra/ring.lean b/library/algebra/ring.lean index 4c51bffaf..0a9d6b947 100644 --- a/library/algebra/ring.lean +++ b/library/algebra/ring.lean @@ -99,7 +99,7 @@ section comm_semiring dvd.elim H₂ (take e, assume H₄ : c = b * e, 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 := dvd.elim H (take c, assume H' : a = 0 * c, H' ⬝ !zero_mul) @@ -117,7 +117,7 @@ section comm_semiring (take d, assume H₁ : b = a * d, 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) := !mul.comm ▸ (dvd_mul_of_dvd_left H _) diff --git a/library/data/list/basic.lean b/library/data/list/basic.lean index 70f31e933..7530a4f3f 100644 --- a/library/data/list/basic.lean +++ b/library/data/list/basic.lean @@ -407,7 +407,7 @@ theorem zip_unzip : ∀ (l : list (A × B)), zip (pr₁ (unzip l)) (pr₂ (unzip have r : zip (pr₁ (unzip l)) (pr₂ (unzip l)) = l, from zip_unzip l, revert r, apply (prod.cases_on (unzip l)), - intros (la, lb, r), + intros [la, lb, r], rewrite -r end diff --git a/library/data/nat/order.lean b/library/data/nat/order.lean index 115955741..df5ecc750 100644 --- a/library/data/nat/order.lean +++ b/library/data/nat/order.lean @@ -105,7 +105,7 @@ theorem lt_add_of_pos_right {n k : ℕ} (H : k > 0) : n < n + k := theorem mul_le_mul_left {n m : ℕ} (H : n ≤ m) (k : ℕ) : k * n ≤ k * m := obtain (l : ℕ) (Hl : n + l = m), from le.elim H, -have H2 : k * n + k * l = k * m, by rewrite ⟨-mul.left_distrib, Hl⟩, +have H2 : k * n + k * l = k * m, by rewrite [-mul.left_distrib, Hl], le.intro H2 theorem mul_le_mul_right {n m : ℕ} (H : n ≤ m) (k : ℕ) : n * k ≤ m * k := diff --git a/library/data/num.lean b/library/data/num.lean index 6ce40d21f..0209b937d 100644 --- a/library/data/num.lean +++ b/library/data/num.lean @@ -498,7 +498,7 @@ namespace pos_num theorem le_trans {a b c : pos_num} : a ≤ b → b ≤ c → a ≤ c := begin - intros (H₁, H₂), + intros [H₁, H₂], rewrite [le_eq_lt_succ at *], apply (@by_cases (a = b)), begin diff --git a/library/data/vector.lean b/library/data/vector.lean index a0e9ab7d1..2231063a8 100644 --- a/library/data/vector.lean +++ b/library/data/vector.lean @@ -125,7 +125,7 @@ namespace vector have H₁ : append t [] == t, from append_nil t, revert H₁, generalize (append t []), rewrite [-add_eq_addl, add_zero], - intros (w, H₁), + intros [w, H₁], rewrite [heq.to_eq H₁], apply heq.refl end diff --git a/library/init/nat.lean b/library/init/nat.lean index 2ce8df1a3..a6cb271a4 100644 --- a/library/init/nat.lean +++ b/library/init/nat.lean @@ -126,7 +126,7 @@ namespace nat definition eq_or_lt_of_le {a b : nat} (H : a ≤ b) : a = b ∨ a < b := begin - cases H with (b, hlt), + cases H with [b, hlt], apply (or.inl rfl), apply (or.inr hlt) end @@ -141,7 +141,7 @@ namespace nat definition le.rec_on {a : nat} {P : nat → Prop} {b : nat} (H : a ≤ b) (H₁ : P a) (H₂ : ∀ b, a < b → P b) : P b := begin - cases H with (b', hlt), + cases H with [b', hlt], apply H₁, apply (H₂ b' hlt) end @@ -190,21 +190,21 @@ namespace nat definition le.trans {a b c : nat} (h₁ : a ≤ b) (h₂ : b ≤ c) : a ≤ c := begin - cases h₁ with (b', hlt), + cases h₁ with [b', hlt], apply h₂, apply (lt.trans hlt h₂) end definition lt_of_le_of_lt {a b c : nat} (h₁ : a ≤ b) (h₂ : b < c) : a < c := begin - cases h₁ with (b', hlt), + cases h₁ with [b', hlt], apply h₂, apply (lt.trans hlt h₂) end definition lt_of_lt_of_le {a b c : nat} (h₁ : a < b) (h₂ : b ≤ c) : a < c := begin - cases h₁ with (b', hlt), + cases h₁ with [b', hlt], apply (lt_of_succ_lt_succ h₂), apply (lt.trans hlt (lt_of_succ_lt_succ h₂)) end diff --git a/src/frontends/lean/parse_rewrite_tactic.cpp b/src/frontends/lean/parse_rewrite_tactic.cpp index e776cef64..28ea1d240 100644 --- a/src/frontends/lean/parse_rewrite_tactic.cpp +++ b/src/frontends/lean/parse_rewrite_tactic.cpp @@ -38,7 +38,7 @@ static expr parse_rule(parser & p, bool use_paren) { static expr parse_rewrite_unfold_core(parser & p) { buffer to_unfold; - if (p.curr_is_token(get_lcurly_tk())) { + if (p.curr_is_token(get_lbracket_tk())) { p.next(); while (true) { to_unfold.push_back(p.check_constant_next("invalid unfold rewrite step, identifier expected")); @@ -46,9 +46,9 @@ static expr parse_rewrite_unfold_core(parser & p) { break; p.next(); } - p.check_token_next(get_rcurly_tk(), "invalid unfold rewrite step, ',' or '}' expected"); + p.check_token_next(get_rbracket_tk(), "invalid unfold rewrite step, ',' or ']' expected"); } else { - to_unfold.push_back(p.check_constant_next("invalid unfold rewrite step, identifier or '{' expected")); + to_unfold.push_back(p.check_constant_next("invalid unfold rewrite step, identifier or '[' expected")); } location loc = parse_tactic_location(p); return mk_rewrite_unfold(to_list(to_unfold), loc); @@ -122,8 +122,7 @@ static expr parse_rewrite_element(parser & p, bool use_paren) { expr parse_rewrite_tactic(parser & p) { buffer elems; - bool lbraket = p.curr_is_token(get_lbracket_tk()); - if (lbraket || p.curr_is_token(get_langle_tk())) { + if (p.curr_is_token(get_lbracket_tk())) { p.next(); while (true) { auto pos = p.pos(); @@ -132,10 +131,7 @@ expr parse_rewrite_tactic(parser & p) { break; p.next(); } - if (lbraket) - p.check_token_next(get_rbracket_tk(), "invalid rewrite tactic, ']' expected"); - else - p.check_token_next(get_rangle_tk(), "invalid rewrite tactic, '⟩' expected"); + p.check_token_next(get_rbracket_tk(), "invalid rewrite tactic, ',' or ']' expected"); } else { auto pos = p.pos(); elems.push_back(p.save_pos(parse_rewrite_element(p, true), pos)); @@ -148,7 +144,7 @@ expr parse_esimp_tactic(parser & p) { auto pos = p.pos(); if (p.curr_is_token(get_up_tk()) || p.curr_is_token(get_caret_tk())) { elems.push_back(p.save_pos(parse_rewrite_unfold(p), pos)); - } else if (p.curr_is_token(get_lcurly_tk())) { + } else if (p.curr_is_token(get_lbracket_tk())) { elems.push_back(p.save_pos(parse_rewrite_unfold_core(p), pos)); } else { location loc = parse_tactic_location(p); @@ -160,7 +156,7 @@ expr parse_esimp_tactic(parser & p) { expr parse_fold_tactic(parser & p) { buffer elems; auto pos = p.pos(); - if (p.curr_is_token(get_lcurly_tk())) { + if (p.curr_is_token(get_lbracket_tk())) { p.next(); while (true) { auto pos = p.pos(); @@ -171,7 +167,7 @@ expr parse_fold_tactic(parser & p) { break; p.next(); } - p.check_token_next(get_rcurly_tk(), "invalid 'fold' tactic, '}' expected"); + p.check_token_next(get_rbracket_tk(), "invalid 'fold' tactic, ',' or ']' expected"); } else { expr e = p.parse_expr(); location loc = parse_tactic_location(p); diff --git a/src/frontends/lean/parser.cpp b/src/frontends/lean/parser.cpp index 8a3c87a41..b2e6eb0cd 100644 --- a/src/frontends/lean/parser.cpp +++ b/src/frontends/lean/parser.cpp @@ -1346,15 +1346,15 @@ optional parser::is_tactic_command(name & id) { expr parser::parse_tactic_expr_list() { auto p = pos(); - check_token_next(get_lparen_tk(), "invalid tactic, '(' expected"); + check_token_next(get_lbracket_tk(), "invalid tactic, '[' expected"); buffer args; - while (!curr_is_token(get_rparen_tk())) { + while (true) { args.push_back(parse_expr()); if (!curr_is_token(get_comma_tk())) break; next(); } - check_token_next(get_rparen_tk(), "invalid tactic, ',' or ')' expected"); + check_token_next(get_rbracket_tk(), "invalid tactic, ',' or ']' expected"); unsigned i = args.size(); expr r = save_pos(mk_constant(get_tactic_expr_list_nil_name()), p); while (i > 0) { @@ -1365,7 +1365,7 @@ expr parser::parse_tactic_expr_list() { } expr parser::parse_tactic_opt_expr_list() { - if (curr_is_token(get_lparen_tk())) { + if (curr_is_token(get_lbracket_tk())) { return parse_tactic_expr_list(); } else if (curr_is_token(get_with_tk())) { next(); diff --git a/tests/lean/487.hlean b/tests/lean/487.hlean index 10e89c4a9..8f97f0496 100644 --- a/tests/lean/487.hlean +++ b/tests/lean/487.hlean @@ -15,5 +15,5 @@ definition foo begin fapply is_retraction.mk, {exact (@ap B A g b b') }, - {intro p, cases p, esimp {eq.ap, eq.rec_on, eq.idp} } + {intro p, cases p, esimp [eq.ap, eq.rec_on, eq.idp] } end diff --git a/tests/lean/hott/433.hlean b/tests/lean/hott/433.hlean index 3d9d2f541..63a2fbfe1 100644 --- a/tests/lean/hott/433.hlean +++ b/tests/lean/hott/433.hlean @@ -104,7 +104,7 @@ namespace pi apply (adjointify (functor_pi f0 f1) (functor_pi (f0⁻¹) (λ(a : A) (b' : B' (f0⁻¹ a)), transport B (retr f0 a) ((f1 (f0⁻¹ a))⁻¹ b')))), intro h, apply eq_of_homotopy, - esimp {functor_pi, function.compose}, -- simplify (and unfold function_pi and function.compose) + esimp [functor_pi, function.compose], -- simplify (and unfold function_pi and function.compose) --first subgoal intro a', esimp, rewrite adj, @@ -115,7 +115,7 @@ namespace pi intro h, beta, apply eq_of_homotopy, intro a, esimp, apply (transport_V (λx, retr f0 a ▹ x = h a) (sect (f1 _) _)), - esimp {function.id}, + esimp [function.id], apply apD end diff --git a/tests/lean/hott/cases.hlean b/tests/lean/hott/cases.hlean index 7e2b52688..b60fa15db 100644 --- a/tests/lean/hott/cases.hlean +++ b/tests/lean/hott/cases.hlean @@ -12,7 +12,7 @@ namespace vec protected definition destruct (v : vec A (succ n)) {P : Π {n : nat}, vec A (succ n) → Type} (H : Π {n : nat} (h : A) (t : vec A n), P (h :: t)) : P v := begin - cases v with (n', h', t'), + cases v with [n', h', t'], apply (H h' t') end diff --git a/tests/lean/hott/inv_bug.hlean b/tests/lean/hott/inv_bug.hlean index b154d50b3..aa0f1db17 100644 --- a/tests/lean/hott/inv_bug.hlean +++ b/tests/lean/hott/inv_bug.hlean @@ -10,15 +10,15 @@ with odd : nat → Type := example : even 1 → empty := begin intro He1, - cases He1 with (a, Ho0), + cases He1 with [a, Ho0], cases Ho0 end example : even 3 → empty := begin intro He3, - cases He3 with (a, Ho2), - cases Ho2 with (a, He1), - cases He1 with (a, Ho0), + cases He3 with [a, Ho2], + cases Ho2 with [a, He1], + cases He1 with [a, Ho0], cases Ho0 end diff --git a/tests/lean/hott/rw_eta.hlean b/tests/lean/hott/rw_eta.hlean index a733af89c..152b85fb5 100644 --- a/tests/lean/hott/rw_eta.hlean +++ b/tests/lean/hott/rw_eta.hlean @@ -25,7 +25,7 @@ local attribute eq_of_homotopy [reducible] definition foo (f g : Πa b, C a b) (H : f ∼2 g) (a : A) : apD100 (eq_of_homotopy2 H) a = H a := begin - esimp {apD100, eq_of_homotopy2, eq_of_homotopy}, + esimp [apD100, eq_of_homotopy2, eq_of_homotopy], rewrite (retr apD10 (λ(a : A), eq_of_homotopy (H a))), apply (retr apD10) end diff --git a/tests/lean/interactive/proof_state_info.input b/tests/lean/interactive/proof_state_info.input index 600bd76f6..0ede4d21d 100644 --- a/tests/lean/interactive/proof_state_info.input +++ b/tests/lean/interactive/proof_state_info.input @@ -5,7 +5,7 @@ open tactic theorem tst (a b c : Prop) : a → b → a ∧ b := begin info, - intros (Ha, Hb), + intros [Ha, Hb], info, apply and.intro, apply Ha, diff --git a/tests/lean/interactive/proof_state_info2.input b/tests/lean/interactive/proof_state_info2.input index 0f899a68f..26ad3b06c 100644 --- a/tests/lean/interactive/proof_state_info2.input +++ b/tests/lean/interactive/proof_state_info2.input @@ -4,7 +4,7 @@ import logic open tactic theorem tst (a b c : Prop) : a → b → a ∧ b := begin - intros (Ha, Hb), + intros [Ha, Hb], apply and.intro, apply Ha, apply Hb, diff --git a/tests/lean/interactive/proof_state_info3.input b/tests/lean/interactive/proof_state_info3.input index b4af4f92d..ffe7fe42e 100644 --- a/tests/lean/interactive/proof_state_info3.input +++ b/tests/lean/interactive/proof_state_info3.input @@ -5,7 +5,7 @@ SYNC 11 theorem tst (a b c : Prop) : a → b → a ∧ b := begin info, - intros (Ha, Hb), + intros [Ha, Hb], info, apply and.intro, apply Ha, diff --git a/tests/lean/interactive/proof_state_info3.input.expected.out b/tests/lean/interactive/proof_state_info3.input.expected.out index 07524c2c6..1743d0ebf 100644 --- a/tests/lean/interactive/proof_state_info3.input.expected.out +++ b/tests/lean/interactive/proof_state_info3.input.expected.out @@ -6,7 +6,7 @@ theorem tst (a b c : Prop) : a → b → a ∧ b := begin info, - intros (Ha, Hb), + intros [Ha, Hb], state, apply and.intro, apply Ha, diff --git a/tests/lean/run/all_goals.lean b/tests/lean/run/all_goals.lean index 54b9df9be..af985d809 100644 --- a/tests/lean/run/all_goals.lean +++ b/tests/lean/run/all_goals.lean @@ -7,6 +7,6 @@ attribute nat.rec_on [unfold-c 2] example (a b c : nat) : a + 0 = 0 + a ∧ b + 0 = 0 + b := begin apply and.intro, - all_goals esimp{of_num}, + all_goals esimp[of_num], all_goals (state; rewrite zero_add) end diff --git a/tests/lean/run/all_goals2.lean b/tests/lean/run/all_goals2.lean index bbbf02b60..eca448e6f 100644 --- a/tests/lean/run/all_goals2.lean +++ b/tests/lean/run/all_goals2.lean @@ -15,6 +15,6 @@ example (a b c : nat) : (a + 0 = 0 + a ∧ b + 0 = 0 + b) ∧ c = c := begin apply and.intro, - apply and.intro ;; (esimp{of_num}; state; rewrite zero_add), + apply and.intro ;; (esimp[of_num]; state; rewrite zero_add), apply rfl end diff --git a/tests/lean/run/beginend3.lean b/tests/lean/run/beginend3.lean index 184857aae..e25e4422a 100644 --- a/tests/lean/run/beginend3.lean +++ b/tests/lean/run/beginend3.lean @@ -3,7 +3,7 @@ open tactic theorem foo (A : Type) (a b c : A) : a = b → b = c → a = c ∧ c = a := begin - intros (Hab, Hbc), + intros [Hab, Hbc], apply and.intro, apply eq.trans, rotate 2, diff --git a/tests/lean/run/clear_tac.lean b/tests/lean/run/clear_tac.lean index 7f9bcee8b..e54e8f36a 100644 --- a/tests/lean/run/clear_tac.lean +++ b/tests/lean/run/clear_tac.lean @@ -2,7 +2,7 @@ import logic example {a b c : Prop} : a → b → c → a ∧ b := begin - intros (Ha, Hb, Hc), + intros [Ha, Hb, Hc], clear Hc, clear c, apply (and.intro Ha Hb), @@ -10,7 +10,7 @@ end example {a b c : Prop} : a → b → c → c ∧ b := begin - intros (Ha, Hb, Hc), + intros [Ha, Hb, Hc], clear Ha, clear a, apply (and.intro Hc Hb), diff --git a/tests/lean/run/clears_tac.lean b/tests/lean/run/clears_tac.lean index 66029dae2..2f5e6e4b0 100644 --- a/tests/lean/run/clears_tac.lean +++ b/tests/lean/run/clears_tac.lean @@ -2,14 +2,14 @@ import logic example {a b c : Prop} : a → b → c → a ∧ b := begin - intros (Ha, Hb, Hc), - clears (Hc, c), + intros [Ha, Hb, Hc], + clears [Hc, c], apply (and.intro Ha Hb), end example {a b c : Prop} : a → b → c → c ∧ b := begin - intros (Ha, Hb, Hc), - clears (Ha, a), + intros [Ha, Hb, Hc], + clears [Ha, a], apply (and.intro Hc Hb), end diff --git a/tests/lean/run/fold.lean b/tests/lean/run/fold.lean index 061d9a2f2..07303c04c 100644 --- a/tests/lean/run/fold.lean +++ b/tests/lean/run/fold.lean @@ -3,7 +3,7 @@ definition id {A : Type} (a : A) := a example (a b c : nat) : id a = id b → a = b := begin intro H, - fold {id a, id b}, + fold [id a, id b], assumption end diff --git a/tests/lean/run/generalizes.lean b/tests/lean/run/generalizes.lean index cb4b8d8e2..0bd726ac6 100644 --- a/tests/lean/run/generalizes.lean +++ b/tests/lean/run/generalizes.lean @@ -2,7 +2,7 @@ import logic theorem tst (A B : Type) (a : A) (b : B) : a == b → b == a := begin - generalizes (a, b, B), - intros (B', b, a, H), + generalizes [a, b, B], + intros [B', b, a, H], apply (heq.symm H), end diff --git a/tests/lean/run/intros.lean b/tests/lean/run/intros.lean index ccd15ba53..9c977cbc9 100644 --- a/tests/lean/run/intros.lean +++ b/tests/lean/run/intros.lean @@ -18,7 +18,7 @@ end theorem tst4 (a b : Prop) : a → b → a ∧ b := begin - intros (Ha, Hb), + intros [Ha, Hb], rapply and.intro, apply Hb, apply Ha diff --git a/tests/lean/run/inv_bug.lean b/tests/lean/run/inv_bug.lean index 905b63674..be3cd1f49 100644 --- a/tests/lean/run/inv_bug.lean +++ b/tests/lean/run/inv_bug.lean @@ -10,15 +10,15 @@ with odd : nat → Prop := example : even 1 → false := begin intro He1, - cases He1 with (a, Ho0), + cases He1 with [a, Ho0], cases Ho0 end example : even 3 → false := begin intro He3, - cases He3 with (a, Ho2), - cases Ho2 with (a, He1), - cases He1 with (a, Ho0), + cases He3 with [a, Ho2], + cases Ho2 with [a, He1], + cases He1 with [a, Ho0], cases Ho0 end diff --git a/tests/lean/run/inv_bug2.lean b/tests/lean/run/inv_bug2.lean index 24a891e34..400c13265 100644 --- a/tests/lean/run/inv_bug2.lean +++ b/tests/lean/run/inv_bug2.lean @@ -7,7 +7,7 @@ namespace vector protected definition destruct2 (v : vector A (succ (succ n))) {P : Π {n : nat}, vector A (succ n) → Type} (H : Π {n : nat} (h : A) (t : vector A n), P (h :: t)) : P v := begin - cases v with (n', h', t'), + cases v with [n', h', t'], apply (H h' t') end end vector diff --git a/tests/lean/run/local_eqns2.lean b/tests/lean/run/local_eqns2.lean index 89bbaa710..badd54c20 100644 --- a/tests/lean/run/local_eqns2.lean +++ b/tests/lean/run/local_eqns2.lean @@ -7,7 +7,7 @@ definition nz_cases_on {C : Π n, fin (succ n) → Type} {n : nat} (f : fin (succ n)) : C n f := begin - reverts (n, f), + reverts [n, f], show ∀ (n : nat) (f : fin (succ n)), C n f | m (fz m) := by apply H₁ | m (fs f') := by apply H₂ diff --git a/tests/lean/run/match_tac4.lean b/tests/lean/run/match_tac4.lean index b7e9ec153..6ec665061 100644 --- a/tests/lean/run/match_tac4.lean +++ b/tests/lean/run/match_tac4.lean @@ -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/nested_begin.lean b/tests/lean/run/nested_begin.lean index 44533be5d..423d8ba02 100644 --- a/tests/lean/run/nested_begin.lean +++ b/tests/lean/run/nested_begin.lean @@ -16,7 +16,7 @@ namespace vector apply (vector.cases_on v₁), exact (assume h : P, h), - intros (n, a, v, h), + intros [n, a, v, h], apply (h rfl), repeat (apply rfl), repeat (apply heq.refl) diff --git a/tests/lean/run/parity.lean b/tests/lean/run/parity.lean index fb4bd43b5..76fb68b17 100644 --- a/tests/lean/run/parity.lean +++ b/tests/lean/run/parity.lean @@ -12,7 +12,7 @@ definition parity : Π (n : nat), Parity n | parity (n+1) := begin have aux : Parity n, from parity n, - cases aux with (k, k), + cases aux with [k, k], begin apply (odd k) end, diff --git a/tests/lean/run/revert_tac.lean b/tests/lean/run/revert_tac.lean index e873457e1..47e6336d8 100644 --- a/tests/lean/run/revert_tac.lean +++ b/tests/lean/run/revert_tac.lean @@ -2,7 +2,7 @@ import logic theorem tst {a b c : Prop} : a → b → c → a ∧ b := begin - intros (Ha, Hb, Hc), + intros [Ha, Hb, Hc], revert Ha, intro Ha2, apply (and.intro Ha2 Hb), @@ -10,7 +10,7 @@ end theorem foo1 {A : Type} (a b c : A) (P : A → Prop) : P a → a = b → P b := begin - intros (Hp, Heq), + intros [Hp, Heq], revert Hp, apply (eq.rec_on Heq), intro Hpa, @@ -19,7 +19,7 @@ end theorem foo2 {A : Type} (a b c : A) (P : A → Prop) : P a → a = b → P b := begin - intros (Hp, Heq), + intros [Hp, Heq], apply (eq.rec_on Heq Hp) end diff --git a/tests/lean/run/reverts_tac.lean b/tests/lean/run/reverts_tac.lean index 6eae43326..9b3e0a34a 100644 --- a/tests/lean/run/reverts_tac.lean +++ b/tests/lean/run/reverts_tac.lean @@ -2,16 +2,16 @@ import logic theorem tst {a b c : Prop} : a → b → c → a ∧ b := begin - intros (Ha, Hb, Hc), - reverts (Hb, Ha), - intros (Hb2, Ha2), + intros [Ha, Hb, Hc], + reverts [Hb, Ha], + intros [Hb2, Ha2], apply (and.intro Ha2 Hb2), end theorem foo1 {A : Type} (a b c : A) (P : A → Prop) : P a → a = b → P b := begin - intros (Hp, Heq), - reverts (Heq, Hp), + intros [Hp, Heq], + reverts [Heq, Hp], intro Heq, apply (eq.rec_on Heq), intro Pa, @@ -20,7 +20,7 @@ end theorem foo2 {A : Type} (a b c : A) (P : A → Prop) : P a → a = b → P b := begin - intros (Hp, Heq), + intros [Hp, Heq], apply (eq.rec_on Heq Hp) end diff --git a/tests/lean/run/rewrite10.lean b/tests/lean/run/rewrite10.lean index 38597d130..a6b0e7028 100644 --- a/tests/lean/run/rewrite10.lean +++ b/tests/lean/run/rewrite10.lean @@ -5,11 +5,11 @@ section variables (a b c d e : nat) theorem T (H1 : a = b) (H2 : b = c + 1) (H3 : c = d) (H4 : e = 1 + d) : a = e := - by rewrite ⟨H1, H2, H3, add.comm, -H4⟩ + by rewrite [H1, H2, H3, add.comm, -H4] end example (x y : ℕ) : (x + y) * (x + y) = x * x + y * x + x * y + y * y := -by rewrite ⟨*mul.left_distrib, *mul.right_distrib, -add.assoc⟩ +by rewrite [*mul.left_distrib, *mul.right_distrib, -add.assoc] definition even (a : nat) := ∃b, a = 2*b @@ -18,10 +18,10 @@ exists.elim H1 (fun (w1 : nat) (Hw1 : a = 2*w1), exists.elim H2 (fun (w2 : nat) (Hw2 : b = 2*w2), exists.intro (w1 + w2) begin - rewrite ⟨Hw1, Hw2, mul.left_distrib⟩ + rewrite [Hw1, Hw2, mul.left_distrib] end)) theorem T2 (a b c : nat) (H1 : a = b) (H2 : b = c + 1) : a ≠ 0 := calc - a = succ c : by rewrite ⟨H1, H2, add_one⟩ + a = succ c : by rewrite [H1, H2, add_one] ... ≠ 0 : succ_ne_zero c diff --git a/tests/lean/run/rewrite12.lean b/tests/lean/run/rewrite12.lean index 043bf30d4..45dacfeb5 100644 --- a/tests/lean/run/rewrite12.lean +++ b/tests/lean/run/rewrite12.lean @@ -3,7 +3,7 @@ open nat variables (f : nat → nat → nat → nat) (a b c : nat) example (H₁ : a = b) (H₂ : f b a b = 0) : f a a a = 0 := -by rewrite ⟨H₁ at -{2}, H₂⟩ +by rewrite [H₁ at -{2}, H₂] example (H₁ : a = b) (H₂ : f b a b = 0) (H₃ : c = f a a a) : c = 0 := -by rewrite ⟨H₁ at H₃ -{2}, H₂ at H₃, H₃⟩ +by rewrite [H₁ at H₃ -{2}, H₂ at H₃, H₃] diff --git a/tests/lean/run/rewrite4.lean b/tests/lean/run/rewrite4.lean index 01dad3705..736fa6a0d 100644 --- a/tests/lean/run/rewrite4.lean +++ b/tests/lean/run/rewrite4.lean @@ -4,10 +4,10 @@ open algebra constant f {A : Type} : A → A → A theorem test1 {A : Type} [s : comm_ring A] (a b c : A) (H : a + 0 = 0) : f a a = f 0 0 := -by rewrite ⟨add_zero at H, H⟩ +by rewrite [add_zero at H, H] theorem test2 {A : Type} [s : comm_ring A] (a b c : A) (H : a + 0 = 0) : f a a = f 0 0 := -by rewrite ⟨add_zero at *, H⟩ +by rewrite [add_zero at *, H] theorem test3 {A : Type} [s : comm_ring A] (a b c : A) (H : a + 0 = 0 + 0) : f a a = f 0 0 := -by rewrite ⟨add_zero at H, zero_add at H, H⟩ +by rewrite [add_zero at H, zero_add at H, H] diff --git a/tests/lean/run/rewrite5.lean b/tests/lean/run/rewrite5.lean index 826274e78..acfb30b7e 100644 --- a/tests/lean/run/rewrite5.lean +++ b/tests/lean/run/rewrite5.lean @@ -6,4 +6,4 @@ variable [s : group A] include s theorem mul.right_inv (a : A) : a * a⁻¹ = 1 := -by rewrite ⟨-{a}inv_inv at {1}, mul.left_inv⟩ +by rewrite [-{a}inv_inv at {1}, mul.left_inv] diff --git a/tests/lean/run/rewriter12.lean b/tests/lean/run/rewriter12.lean index fb4955b56..e20076f20 100644 --- a/tests/lean/run/rewriter12.lean +++ b/tests/lean/run/rewriter12.lean @@ -3,7 +3,7 @@ open nat constant f : nat → nat example (x y : nat) (H1 : (λ z, z + 0) x = y) : f x = f y := -by rewrite [▸* at H1, ^{add, nat.rec_on, of_num} at H1, H1] +by rewrite [▸* at H1, ^[add, nat.rec_on, of_num] at H1, H1] example (x y : nat) (H1 : (λ z, z + 0) x = y) : f x = f y := -by rewrite [▸* at H1, ↑{add, nat.rec_on, of_num} at H1, H1] +by rewrite [▸* at H1, ↑[add, nat.rec_on, of_num] at H1, H1] diff --git a/tests/lean/run/rewriter6.lean b/tests/lean/run/rewriter6.lean index b5d3e7bd5..75c3277d6 100644 --- a/tests/lean/run/rewriter6.lean +++ b/tests/lean/run/rewriter6.lean @@ -10,4 +10,4 @@ by rewrite ^sub -- unfold sub definition double (x : int) := x + x theorem double_zero (x : int) : double (0 + x) = (1 + 1)*x := -by rewrite ⟨↑double, zero_add, mul.right_distrib, one_mul⟩ +by rewrite [↑double, zero_add, mul.right_distrib, one_mul] diff --git a/tests/lean/run/rewriter7.lean b/tests/lean/run/rewriter7.lean index 8a9787e43..fa00de667 100644 --- a/tests/lean/run/rewriter7.lean +++ b/tests/lean/run/rewriter7.lean @@ -6,4 +6,4 @@ constant f : int → int definition double (x : int) := x + x theorem tst1 (x y : int) (H1 : double x = 0) (H2 : double y = 0) (H3 : f (double y) = 0) (H4 : y > 0) : f (x + x) = 0 := -by rewrite ⟨↑double at H1, H1, H2 at H3, H3⟩ +by rewrite [↑double at H1, H1, H2 at H3, H3] diff --git a/tests/lean/run/tactic_op_overload_bug.lean b/tests/lean/run/tactic_op_overload_bug.lean index 2a2e15307..79f8b250b 100644 --- a/tests/lean/run/tactic_op_overload_bug.lean +++ b/tests/lean/run/tactic_op_overload_bug.lean @@ -8,7 +8,7 @@ definition lt_pred (a b : pos_num) : Prop := lt a b = tt definition not_lt_one1 (a : pos_num) : ¬ lt_pred a one := begin - esimp {lt_pred}, + esimp [lt_pred], intro H, apply (absurd_of_eq_ff_of_eq_tt (lt_one_right_eq_ff a) H) end @@ -19,7 +19,7 @@ print raw intro -- intro is overloaded definition not_lt_one2 (a : pos_num) : ¬ lt_pred a one := begin - esimp {lt_pred}, + esimp [lt_pred], intro H, apply (absurd_of_eq_ff_of_eq_tt (lt_one_right_eq_ff a) H) end diff --git a/tests/lean/run/vec_inv2.lean b/tests/lean/run/vec_inv2.lean index be5cacfb8..35c6426ed 100644 --- a/tests/lean/run/vec_inv2.lean +++ b/tests/lean/run/vec_inv2.lean @@ -22,7 +22,7 @@ namespace vector protected definition destruct (v : vector A (succ n)) {P : Π {n : nat}, vector A (succ n) → Type} (H : Π {n : nat} (h : A) (t : vector A n), P (cons h t)) : P v := begin - cases v with (h', n', t'), + cases v with [h', n', t'], apply (H h' t') end diff --git a/tests/lean/run/vec_inv3.lean b/tests/lean/run/vec_inv3.lean index e92290a16..7fa820c1d 100644 --- a/tests/lean/run/vec_inv3.lean +++ b/tests/lean/run/vec_inv3.lean @@ -18,6 +18,6 @@ namespace vector protected definition destruct (v : vector A (succ n)) {P : Π {n : nat}, vector A (succ n) → Type} (H : Π {n : nat} (h : A) (t : vector A n), P (cons h t)) : P v := - by cases v with (h', n', t'); apply (H h' t') + by cases v with [h', n', t']; apply (H h' t') end vector diff --git a/tests/lean/run/vector.lean b/tests/lean/run/vector.lean index ad84cbfe9..9ac6f64e3 100644 --- a/tests/lean/run/vector.lean +++ b/tests/lean/run/vector.lean @@ -107,13 +107,13 @@ namespace vector @vector.brec_on A P n w (λ (n : nat) (w : vector A n), begin - cases w with (n₁, h₁, t₁), + cases w with [n₁, h₁, t₁], show @below A P zero vnil → vector B zero → vector C zero, from λ b v, vnil, show @below A P (succ n₁) (h₁ :: t₁) → vector B (succ n₁) → vector C (succ n₁), from λ b v, begin - cases v with (n₂, h₂, t₂), + cases v with [n₂, h₂, t₂], have r : vector B n₂ → vector C n₂, from pr₁ b, exact ((f h₁ h₂) :: r t₂), end