feat(frontends/lean): remove difference between 'have' and 'assert'

This commit is contained in:
Leonardo de Moura 2016-02-29 11:28:20 -08:00
parent 23079a75a7
commit 101cf1ec4c
25 changed files with 61 additions and 70 deletions

View file

@ -345,13 +345,13 @@ section
theorem eq_of_mul_eq_mul_right {a b c : A} (Ha : a ≠ 0) (H : b * a = c * a) : b = c := theorem eq_of_mul_eq_mul_right {a b c : A} (Ha : a ≠ 0) (H : b * a = c * a) : b = c :=
have b * a - c * a = 0, from iff.mp !eq_iff_sub_eq_zero H, have b * a - c * a = 0, from iff.mp !eq_iff_sub_eq_zero H,
have (b - c) * a = 0, using this, by rewrite [mul_sub_right_distrib, this], have (b - c) * a = 0, by rewrite [mul_sub_right_distrib, this],
have b - c = 0, from sum_resolve_left (eq_zero_sum_eq_zero_of_mul_eq_zero this) Ha, have b - c = 0, from sum_resolve_left (eq_zero_sum_eq_zero_of_mul_eq_zero this) Ha,
iff.elim_right !eq_iff_sub_eq_zero this iff.elim_right !eq_iff_sub_eq_zero this
theorem eq_of_mul_eq_mul_left {a b c : A} (Ha : a ≠ 0) (H : a * b = a * c) : b = c := theorem eq_of_mul_eq_mul_left {a b c : A} (Ha : a ≠ 0) (H : a * b = a * c) : b = c :=
have a * b - a * c = 0, from iff.mp !eq_iff_sub_eq_zero H, have a * b - a * c = 0, from iff.mp !eq_iff_sub_eq_zero H,
have a * (b - c) = 0, using this, by rewrite [mul_sub_left_distrib, this], have a * (b - c) = 0, by rewrite [mul_sub_left_distrib, this],
have b - c = 0, from sum_resolve_right (eq_zero_sum_eq_zero_of_mul_eq_zero this) Ha, have b - c = 0, from sum_resolve_right (eq_zero_sum_eq_zero_of_mul_eq_zero this) Ha,
iff.elim_right !eq_iff_sub_eq_zero this iff.elim_right !eq_iff_sub_eq_zero this

View file

@ -347,13 +347,13 @@ section
theorem eq_of_mul_eq_mul_right {a b c : A} (Ha : a ≠ 0) (H : b * a = c * a) : b = c := theorem eq_of_mul_eq_mul_right {a b c : A} (Ha : a ≠ 0) (H : b * a = c * a) : b = c :=
have b * a - c * a = 0, from iff.mp !eq_iff_sub_eq_zero H, have b * a - c * a = 0, from iff.mp !eq_iff_sub_eq_zero H,
have (b - c) * a = 0, using this, by rewrite [mul_sub_right_distrib, this], have (b - c) * a = 0, by rewrite [mul_sub_right_distrib, this],
have b - c = 0, from or_resolve_left (eq_zero_or_eq_zero_of_mul_eq_zero this) Ha, have b - c = 0, from or_resolve_left (eq_zero_or_eq_zero_of_mul_eq_zero this) Ha,
iff.elim_right !eq_iff_sub_eq_zero this iff.elim_right !eq_iff_sub_eq_zero this
theorem eq_of_mul_eq_mul_left {a b c : A} (Ha : a ≠ 0) (H : a * b = a * c) : b = c := theorem eq_of_mul_eq_mul_left {a b c : A} (Ha : a ≠ 0) (H : a * b = a * c) : b = c :=
have a * b - a * c = 0, from iff.mp !eq_iff_sub_eq_zero H, have a * b - a * c = 0, from iff.mp !eq_iff_sub_eq_zero H,
have a * (b - c) = 0, using this, by rewrite [mul_sub_left_distrib, this], have a * (b - c) = 0, by rewrite [mul_sub_left_distrib, this],
have b - c = 0, from or_resolve_right (eq_zero_or_eq_zero_of_mul_eq_zero this) Ha, have b - c = 0, from or_resolve_right (eq_zero_or_eq_zero_of_mul_eq_zero this) Ha,
iff.elim_right !eq_iff_sub_eq_zero this iff.elim_right !eq_iff_sub_eq_zero this

View file

@ -259,9 +259,9 @@ namespace vector
| (n+1) (m+1) (x::xs) (y::ys) h₁ h₂ := | (n+1) (m+1) (x::xs) (y::ys) h₁ h₂ :=
assert e₁ : n = m, from succ.inj h₂, assert e₁ : n = m, from succ.inj h₂,
assert e₂ : x = y, begin unfold to_list at h₁, injection h₁, assumption end, assert e₂ : x = y, begin unfold to_list at h₁, injection h₁, assumption end,
have to_list xs = to_list ys, begin unfold to_list at h₁, injection h₁, assumption end, have e₃ : to_list xs = to_list ys, begin unfold to_list at h₁, injection h₁, assumption end,
assert xs == ys, from heq_of_list_eq this e₁, assert xs == ys, from heq_of_list_eq e₃ e₁,
assert y :: xs == y :: ys, begin clear heq_of_list_eq h₁ h₂, revert xs ys this, induction e₁, intro xs ys h, rewrite [eq_of_heq h] end, assert y :: xs == y :: ys, begin clear heq_of_list_eq h₁ h₂ e₃, revert xs ys this, induction e₁, intro xs ys h, rewrite [eq_of_heq h] end,
show x :: xs == y :: ys, by rewrite e₂; exact this show x :: xs == y :: ys, by rewrite e₂; exact this
theorem list_eq_of_heq {n m} {v₁ : vector A n} {v₂ : vector A m} : v₁ == v₂ → n = m → to_list v₁ = to_list v₂ := theorem list_eq_of_heq {n m} {v₁ : vector A n} {v₂ : vector A m} : v₁ == v₂ → n = m → to_list v₁ = to_list v₂ :=

View file

@ -753,7 +753,7 @@ assume h, perm.induction_on h
have nodup l₁, from nodup_of_nodup_cons this, have nodup l₁, from nodup_of_nodup_cons this,
have x ∉ l₁, from not_mem_of_nodup_cons `nodup (x::l₁)`, have x ∉ l₁, from not_mem_of_nodup_cons `nodup (x::l₁)`,
have y ∉ x::l₁, from not_mem_of_nodup_cons nd, have y ∉ x::l₁, from not_mem_of_nodup_cons nd,
have x ≠ y, using this, from suppose x = y, begin subst x, exact absurd !mem_cons `y ∉ y::l₁` end, have x ≠ y, from suppose x = y, begin subst x, exact absurd !mem_cons `y ∉ y::l₁` end,
have y ∉ l₁, from not_mem_of_not_mem_cons `y ∉ x::l₁`, have y ∉ l₁, from not_mem_of_not_mem_cons `y ∉ x::l₁`,
have x ∉ y::l₁, from not_mem_cons_of_ne_of_not_mem `x ≠ y` `x ∉ l₁`, have x ∉ y::l₁, from not_mem_cons_of_ne_of_not_mem `x ≠ y` `x ∉ l₁`,
have nodup (y::l₁), from nodup_cons `y ∉ l₁` `nodup l₁`, have nodup (y::l₁), from nodup_cons `y ∉ l₁` `nodup l₁`,

View file

@ -61,7 +61,7 @@ begin
suppose x ∈ upto n, suppose x ∈ upto n,
have x < n, from lt_of_mem_upto this, have x < n, from lt_of_mem_upto this,
suppose x ∈ '{n}, suppose x ∈ '{n},
have x = n, using this, by rewrite -mem_singleton_iff; apply this, have x = n, by rewrite -mem_singleton_iff; apply this,
have n < n, from eq.subst this `x < n`, have n < n, from eq.subst this `x < n`,
show false, from !lt.irrefl this), show false, from !lt.irrefl this),
rewrite [sum_up_to_succ, ih, upto_succ, Sum_union _ H, Sum_singleton] rewrite [sum_up_to_succ, ih, upto_succ, Sum_union _ H, Sum_singleton]
@ -152,7 +152,7 @@ begin
suppose x ∈ upto n, suppose x ∈ upto n,
have x < n, from lt_of_mem_upto this, have x < n, from lt_of_mem_upto this,
suppose x ∈ '{n}, suppose x ∈ '{n},
have x = n, using this, by rewrite -mem_singleton_iff; apply this, have x = n, by rewrite -mem_singleton_iff; apply this,
have n < n, from eq.subst this `x < n`, have n < n, from eq.subst this `x < n`,
show false, from !lt.irrefl this), show false, from !lt.irrefl this),
rewrite [prod_up_to_succ, ih, upto_succ, Prod_union _ H, Prod_singleton] rewrite [prod_up_to_succ, ih, upto_succ, Prod_union _ H, Prod_singleton]

View file

@ -750,7 +750,7 @@ ext (take x, iff.intro
(suppose x ∈ sUnion '{s}, (suppose x ∈ sUnion '{s},
obtain u [(Hu : u ∈ '{s}) (xu : x ∈ u)], from this, obtain u [(Hu : u ∈ '{s}) (xu : x ∈ u)], from this,
have u = s, from eq_of_mem_singleton Hu, have u = s, from eq_of_mem_singleton Hu,
show x ∈ s, using this, by rewrite -this; apply xu) show x ∈ s, by rewrite -this; apply xu)
(suppose x ∈ s, (suppose x ∈ s,
mem_sUnion this (mem_singleton s))) mem_sUnion this (mem_singleton s)))
@ -798,7 +798,7 @@ ext (take x, iff.intro
take t, suppose t ∈ compl ' S, take t, suppose t ∈ compl ' S,
obtain t' [(Ht' : t' ∈ S) (Ht : -t' = t)], from this, obtain t' [(Ht' : t' ∈ S) (Ht : -t' = t)], from this,
have x ∈ -t', from suppose x ∈ t', H (mem_sUnion this Ht'), have x ∈ -t', from suppose x ∈ t', H (mem_sUnion this Ht'),
show x ∈ t, using this, by rewrite -Ht; apply this) show x ∈ t, by rewrite -Ht; apply this)
(assume H : x ∈ ⋂₀ (compl ' S), (assume H : x ∈ ⋂₀ (compl ' S),
suppose x ∈ ⋃₀ S, suppose x ∈ ⋃₀ S,
obtain t [(tS : t ∈ S) (xt : x ∈ t)], from this, obtain t [(tS : t ∈ S) (xt : x ∈ t)], from this,

View file

@ -135,8 +135,7 @@ section
assume H' : eventually (λx, ¬ (P x Q x)) F, assume H' : eventually (λx, ¬ (P x Q x)) F,
have (λx, ¬ (P x Q x)) = λ x, ¬ P x ∧ ¬ Q x, have (λx, ¬ (P x Q x)) = λ x, ¬ P x ∧ ¬ Q x,
by apply funext; intro x; rewrite not_or_iff_not_and_not, by apply funext; intro x; rewrite not_or_iff_not_and_not,
show false, using this, show false, by rewrite this at H'; exact H (eventually_of_eventually_and_left H')
by rewrite this at H'; exact H (eventually_of_eventually_and_left H')
theorem frequently_inr (H : frequently Q F) : frequently (λx, P x Q x) F := theorem frequently_inr (H : frequently Q F) : frequently (λx, P x Q x) F :=
begin begin

View file

@ -37,7 +37,7 @@ metric_space.dist_triangle x y z
proposition dist_nonneg (x y : M) : 0 ≤ dist x y := proposition dist_nonneg (x y : M) : 0 ≤ dist x y :=
have dist x y + dist y x ≥ 0, by rewrite -(dist_self x); apply dist_triangle, have dist x y + dist y x ≥ 0, by rewrite -(dist_self x); apply dist_triangle,
have 2 * dist x y ≥ 0, using this, have 2 * dist x y ≥ 0,
by krewrite [-real.one_add_one, right_distrib, +one_mul, dist_comm at {2}]; apply this, by krewrite [-real.one_add_one, right_distrib, +one_mul, dist_comm at {2}]; apply this,
nonneg_of_mul_nonneg_left this two_pos nonneg_of_mul_nonneg_left this two_pos
@ -135,7 +135,7 @@ exists.intro (N + k)
(take n : , assume nge : n ≥ N + k, (take n : , assume nge : n ≥ N + k,
have n - k ≥ N, from nat.le_sub_of_add_le nge, have n - k ≥ N, from nat.le_sub_of_add_le nge,
have dist (X (n - k + k)) y < ε, from HN (n - k) this, have dist (X (n - k + k)) y < ε, from HN (n - k) this,
show dist (X n) y < ε, using this, show dist (X n) y < ε,
by rewrite [(nat.sub_add_cancel (le.trans !le_add_left nge)) at this]; exact this) by rewrite [(nat.sub_add_cancel (le.trans !le_add_left nge)) at this]; exact this)
proposition converges_to_seq_of_converges_to_seq_offset_left proposition converges_to_seq_of_converges_to_seq_offset_left

View file

@ -214,7 +214,7 @@ obtain N (HN : ∀ n, n ≥ N → norm (X n - 0) < ε), from HX `ε > 0`,
exists.intro N exists.intro N
(take n, assume Hn : n ≥ N, (take n, assume Hn : n ≥ N,
have norm (X n) < ε, begin rewrite -(sub_zero (X n)), apply HN n Hn end, have norm (X n) < ε, begin rewrite -(sub_zero (X n)), apply HN n Hn end,
show abs (norm (X n) - 0) < ε, using this, show abs (norm (X n) - 0) < ε,
by rewrite [sub_zero, abs_of_nonneg !norm_nonneg]; apply this) by rewrite [sub_zero, abs_of_nonneg !norm_nonneg]; apply this)
proposition converges_to_seq_zero_of_norm_converges_to_seq_zero proposition converges_to_seq_zero_of_norm_converges_to_seq_zero
@ -227,7 +227,7 @@ exists.intro (N : )
have HN' : abs (norm (X n) - 0) < ε, from HN n Hn, have HN' : abs (norm (X n) - 0) < ε, from HN n Hn,
have norm (X n) < ε, have norm (X n) < ε,
by+ rewrite [sub_zero at HN', abs_of_nonneg !norm_nonneg at HN']; apply HN', by+ rewrite [sub_zero at HN', abs_of_nonneg !norm_nonneg at HN']; apply HN',
show norm (X n - 0) < ε, using this, show norm (X n - 0) < ε,
by rewrite sub_zero; apply this) by rewrite sub_zero; apply this)
proposition norm_converges_to_seq_zero_iff (X : → V) : proposition norm_converges_to_seq_zero_iff (X : → V) :

View file

@ -140,7 +140,7 @@ have Hb' : ∀ x, x ∈ negX → -b ≤ x,
have HX : X = {x | -x ∈ negX}, have HX : X = {x | -x ∈ negX},
from set.ext (take x, by rewrite [↑set_of, ↑mem, +neg_neg]), from set.ext (take x, by rewrite [↑set_of, ↑mem, +neg_neg]),
show inf {x | -x ∈ X} = - sup X, show inf {x | -x ∈ X} = - sup X,
using HX Hb' nonempty_negX, by rewrite [HX at {2}, sup_neg nonempty_negX Hb', neg_neg] by rewrite [HX at {2}, sup_neg nonempty_negX Hb', neg_neg]
end end
end real end real
@ -453,11 +453,11 @@ have Hi' : ∀ j, j ≥ i → sX - ε < X j, from
exists.intro i exists.intro i
(take j, assume Hj : j ≥ i, (take j, assume Hj : j ≥ i,
have X j - sX ≤ 0, from sub_nonpos_of_le (Xle j), have X j - sX ≤ 0, from sub_nonpos_of_le (Xle j),
have eq₁ : abs (X j - sX) = sX - X j, using this, by rewrite [abs_of_nonpos this, neg_sub], have eq₁ : abs (X j - sX) = sX - X j, by rewrite [abs_of_nonpos this, neg_sub],
have sX - ε < X j, from lt_of_lt_of_le (by rewrite Hi; apply H₂x') (nondecX Hj), have sX - ε < X j, from lt_of_lt_of_le (by rewrite Hi; apply H₂x') (nondecX Hj),
have sX < X j + ε, from lt_add_of_sub_lt_right this, have sX < X j + ε, from lt_add_of_sub_lt_right this,
have sX - X j < ε, from sub_lt_left_of_lt_add this, have sX - X j < ε, from sub_lt_left_of_lt_add this,
show (abs (X j - sX)) < ε, using eq₁ this, by rewrite eq₁; exact this) show (abs (X j - sX)) < ε, by rewrite eq₁; exact this)
definition nonincreasing (X : ) : Prop := ∀ ⦃i j⦄, i ≤ j → X i ≥ X j definition nonincreasing (X : ) : Prop := ∀ ⦃i j⦄, i ≤ j → X i ≥ X j
@ -522,7 +522,6 @@ theorem pow_converges_to_seq_zero {x : } (H : abs x < 1) :
(λ n, x^n) ⟶ 0 in := (λ n, x^n) ⟶ 0 in :=
suffices H' : (λ n, (abs x)^n) ⟶ 0 in , from suffices H' : (λ n, (abs x)^n) ⟶ 0 in , from
have (λ n, (abs x)^n) = (λ n, abs (x^n)), from funext (take n, eq.symm !abs_pow), have (λ n, (abs x)^n) = (λ n, abs (x^n)), from funext (take n, eq.symm !abs_pow),
using this,
by rewrite this at H'; exact converges_to_seq_zero_of_abs_converges_to_seq_zero H', by rewrite this at H'; exact converges_to_seq_zero_of_abs_converges_to_seq_zero H',
let aX := (λ n, (abs x)^n), let aX := (λ n, (abs x)^n),
iaX := real.inf (aX ' univ), iaX := real.inf (aX ' univ),

View file

@ -43,17 +43,17 @@ theorem measurable_cInter {s : → set X} (H : ∀ i, measurable (s i)) :
measurable (⋂ i, s i) := measurable (⋂ i, s i) :=
have ∀ i, measurable (-(s i)), from take i, measurable_compl (H i), have ∀ i, measurable (-(s i)), from take i, measurable_compl (H i),
have measurable (-( i, -(s i))), from measurable_compl (measurable_cUnion this), have measurable (-( i, -(s i))), from measurable_compl (measurable_cUnion this),
show measurable (⋂ i, s i), using this, by rewrite Inter_eq_comp_Union_comp; apply this show measurable (⋂ i, s i), by rewrite Inter_eq_comp_Union_comp; apply this
theorem measurable_union {s t : set X} (Hs : measurable s) (Ht : measurable t) : theorem measurable_union {s t : set X} (Hs : measurable s) (Ht : measurable t) :
measurable (s t) := measurable (s t) :=
have ∀ i, measurable (bin_ext s t i), by intro i; cases i; exact Hs; exact Ht, have ∀ i, measurable (bin_ext s t i), by intro i; cases i; exact Hs; exact Ht,
show measurable (s t), using this, by rewrite -Union_bin_ext; exact measurable_cUnion this show measurable (s t), by rewrite -Union_bin_ext; exact measurable_cUnion this
theorem measurable_inter {s t : set X} (Hs : measurable s) (Ht : measurable t) : theorem measurable_inter {s t : set X} (Hs : measurable s) (Ht : measurable t) :
measurable (s ∩ t) := measurable (s ∩ t) :=
have ∀ i, measurable (bin_ext s t i), by intro i; cases i; exact Hs; exact Ht, have ∀ i, measurable (bin_ext s t i), by intro i; cases i; exact Hs; exact Ht,
show measurable (s ∩ t), using this, by rewrite -Inter_bin_ext; exact measurable_cInter this show measurable (s ∩ t), by rewrite -Inter_bin_ext; exact measurable_cInter this
theorem measurable_diff {s t : set X} (Hs : measurable s) (Ht : measurable t) : theorem measurable_diff {s t : set X} (Hs : measurable s) (Ht : measurable t) :
measurable (s \ t) := measurable (s \ t) :=

View file

@ -100,15 +100,14 @@ section
have bnnz : (b : rat)^n ≠ 0, have bnnz : (b : rat)^n ≠ 0,
from assume bneqz, bnz (_root_.eq_zero_of_pow_eq_zero bneqz), from assume bneqz, bnz (_root_.eq_zero_of_pow_eq_zero bneqz),
have a^n /[rat] (b:rat)^n = c, have a^n /[rat] (b:rat)^n = c,
using bnz, begin rewrite [*of_int_pow, -div_pow, -eq_num_div_denom, -H] end, begin rewrite [*of_int_pow, -div_pow, -eq_num_div_denom, -H] end,
have (a^n : rat) = c * (b:rat)^n, have (a^n : rat) = c * (b:rat)^n,
from eq.symm (!mul_eq_of_eq_div bnnz (eq.symm this)), from eq.symm (!mul_eq_of_eq_div bnnz (eq.symm this)),
have a^n = c * b^n, -- int version have a^n = c * b^n, -- int version
using this, by rewrite [-of_int_pow at this, -of_int_mul at this]; exact of_int.inj this, by rewrite [-of_int_pow at this, -of_int_mul at this]; exact of_int.inj this,
have (abs a)^n = abs c * (abs b)^n, have (abs a)^n = abs c * (abs b)^n,
using this, by rewrite [-abs_pow, this, abs_mul, abs_pow], by rewrite [-abs_pow, this, abs_mul, abs_pow],
have H₁ : (nat_abs a)^n = nat_abs c * (nat_abs b)^n, have H₁ : (nat_abs a)^n = nat_abs c * (nat_abs b)^n,
using this,
begin apply int.of_nat.inj, rewrite [int.of_nat_mul, +int.of_nat_pow, +of_nat_nat_abs], begin apply int.of_nat.inj, rewrite [int.of_nat_mul, +int.of_nat_pow, +of_nat_nat_abs],
exact this end, exact this end,
have H₂ : nat.coprime (nat_abs a) (nat_abs b), have H₂ : nat.coprime (nat_abs a) (nat_abs b),
@ -118,7 +117,7 @@ section
(suppose q = 0, (suppose q = 0,
by rewrite this) by rewrite this)
(suppose qne0 : q ≠ 0, (suppose qne0 : q ≠ 0,
using H₁ H₂, begin begin
have ane0 : a ≠ 0, from have ane0 : a ≠ 0, from
suppose aeq0 : a = 0, suppose aeq0 : a = 0,
have qeq0 : q = 0, have qeq0 : q = 0,
@ -133,13 +132,13 @@ section
show nat_abs b = 1, from (root_irrational npos (pos_of_ne_zero this) H₂ H₁) show nat_abs b = 1, from (root_irrational npos (pos_of_ne_zero this) H₂ H₁)
end), end),
show b = 1, show b = 1,
using this, begin rewrite [-of_nat_nat_abs_of_nonneg (le_of_lt !denom_pos), this] end begin rewrite [-of_nat_nat_abs_of_nonneg (le_of_lt !denom_pos), this] end
theorem eq_num_pow_of_pow_eq {q : } {n : } {c : } (npos : n > 0) (H : q^n = c) : theorem eq_num_pow_of_pow_eq {q : } {n : } {c : } (npos : n > 0) (H : q^n = c) :
c = (num q)^n := c = (num q)^n :=
have denom q = 1, have denom q = 1,
from denom_eq_one_of_pow_eq npos H, from denom_eq_one_of_pow_eq npos H,
have of_int c = of_int ((num q)^n), using this, have of_int c = of_int ((num q)^n),
by rewrite [-H, eq_num_div_denom q at {1}, this, of_int_one, div_one, of_int_pow], by rewrite [-H, eq_num_div_denom q at {1}, this, of_int_one, div_one, of_int_pow],
show c = (num q)^n , from of_int.inj this show c = (num q)^n , from of_int.inj this
end end
@ -159,7 +158,7 @@ section
by note H := (pos_of_prime primep); rewrite this at H; exfalso; exact !lt.irrefl H), by note H := (pos_of_prime primep); rewrite this at H; exfalso; exact !lt.irrefl H),
assert agtz : a > 0, from pos_of_ne_zero assert agtz : a > 0, from pos_of_ne_zero
(suppose a = 0, (suppose a = 0,
show false, using npos pnez, by revert peq; rewrite [this, zero_pow npos]; exact pnez), show false, by revert peq; rewrite [this, zero_pow npos]; exact pnez),
have n * mult p a = 1, from calc have n * mult p a = 1, from calc
n * mult p a = mult p (a^n) : begin rewrite [mult_pow n agtz primep] end n * mult p a = mult p (a^n) : begin rewrite [mult_pow n agtz primep] end
... = mult p p : peq ... = mult p p : peq
@ -168,7 +167,7 @@ section
from dvd_of_mul_right_eq this, from dvd_of_mul_right_eq this,
have n = 1, have n = 1,
from eq_one_of_dvd_one this, from eq_one_of_dvd_one this,
show false, using this, show false,
by rewrite this at ngt1; exact !lt.irrefl ngt1 by rewrite this at ngt1; exact !lt.irrefl ngt1
open int rat open int rat
@ -180,7 +179,7 @@ section
have npos : n > 0, from lt.trans dec_trivial ngt1, have npos : n > 0, from lt.trans dec_trivial ngt1,
suppose q^n = p, suppose q^n = p,
have p = (num q)^n, from eq_num_pow_of_pow_eq npos this, have p = (num q)^n, from eq_num_pow_of_pow_eq npos this,
have p = (nat_abs (num q))^n, using this numq, have p = (nat_abs (num q))^n,
by apply of_nat.inj; rewrite [this, of_nat_pow, of_nat_nat_abs_of_nonneg numq], by apply of_nat.inj; rewrite [this, of_nat_pow, of_nat_nat_abs_of_nonneg numq],
show false, from not_eq_pow_of_prime _ ngt1 primep this show false, from not_eq_pow_of_prime _ ngt1 primep this
end end
@ -204,15 +203,15 @@ section
by rewrite -mul.assoc at H; apply H, by rewrite -mul.assoc at H; apply H,
have a / (gcd a b) = c * b / gcd (c * b) a, have a / (gcd a b) = c * b / gcd (c * b) a,
from div_gcd_eq_div_gcd this bpos apos, from div_gcd_eq_div_gcd this bpos apos,
have a = c * b / gcd c a, using this, have a = c * b / gcd c a,
by revert this; rewrite [↑coprime at co, co, int.div_one, H₁]; intros; assumption, by revert this; rewrite [↑coprime at co, co, int.div_one, H₁]; intros; assumption,
have a = b * (c / gcd c a), using this, have a = b * (c / gcd c a),
by revert this; rewrite [mul.comm, !int.mul_div_assoc !gcd_dvd_left]; intros; assumption, by revert this; rewrite [mul.comm, !int.mul_div_assoc !gcd_dvd_left]; intros; assumption,
have b a, have b a,
from dvd_of_mul_right_eq this⁻¹, from dvd_of_mul_right_eq this⁻¹,
have b gcd a b, have b gcd a b,
from dvd_gcd this !dvd.refl, from dvd_gcd this !dvd.refl,
have b 1, using this, have b 1,
by rewrite [↑coprime at co, co at this]; apply this, by rewrite [↑coprime at co, co at this]; apply this,
show b = 1, show b = 1,
from eq_one_of_dvd_one (le_of_lt bpos) this from eq_one_of_dvd_one (le_of_lt bpos) this

View file

@ -25,7 +25,7 @@ definition Open (s : set X) : Prop := s ∈ opens X
theorem Open_empty : Open (∅ : set X) := theorem Open_empty : Open (∅ : set X) :=
have ∅ ⊆ opens X, from empty_subset _, have ∅ ⊆ opens X, from empty_subset _,
have ⋃₀ ∅ ∈ opens X, from sUnion_mem_opens this, have ⋃₀ ∅ ∈ opens X, from sUnion_mem_opens this,
show ∅ ∈ opens X, using this, by rewrite -sUnion_empty; apply this show ∅ ∈ opens X, by rewrite -sUnion_empty; apply this
theorem Open_univ : Open (univ : set X) := theorem Open_univ : Open (univ : set X) :=
univ_mem_opens X univ_mem_opens X
@ -42,7 +42,7 @@ using this, by rewrite Union_eq_sUnion_image; apply Open_sUnion this
theorem Open_union {s t : set X} (Hs : Open s) (Ht : Open t) : Open (s t) := theorem Open_union {s t : set X} (Hs : Open s) (Ht : Open t) : Open (s t) :=
have ∀ i, Open (bin_ext s t i), by intro i; cases i; exact Hs; exact Ht, have ∀ i, Open (bin_ext s t i), by intro i; cases i; exact Hs; exact Ht,
show Open (s t), using this, by rewrite -Union_bin_ext; exact Open_Union this show Open (s t), by rewrite -Union_bin_ext; exact Open_Union this
theorem Open_inter {s t : set X} (Hs : Open s) (Ht : Open t) : Open (s ∩ t) := theorem Open_inter {s t : set X} (Hs : Open s) (Ht : Open t) : Open (s ∩ t) :=
inter_mem_opens X Hs Ht inter_mem_opens X Hs Ht

View file

@ -404,11 +404,6 @@ static expr parse_proof(parser & p, expr const & prop) {
} else if (p.curr_is_token(get_by_plus_tk())) { } else if (p.curr_is_token(get_by_plus_tk())) {
auto pos = p.pos(); auto pos = p.pos();
return parse_by_plus(p, 0, nullptr, pos); return parse_by_plus(p, 0, nullptr, pos);
} else if (p.curr_is_token(get_using_tk())) {
// parse: 'using' locals* ',' proof
auto using_pos = p.pos();
p.next();
return parse_using_expr(p, prop, using_pos);
} else { } else {
throw parser_error("invalid expression, 'by', 'begin', 'proof', 'using' or 'from' expected", p.pos()); throw parser_error("invalid expression, 'by', 'begin', 'proof', 'using' or 'from' expected", p.pos());
} }
@ -466,6 +461,7 @@ static expr parse_have_core(parser & p, pos_info const & pos, optional<expr> con
} }
p.check_token_next(get_comma_tk(), "invalid 'have/assert' declaration, ',' expected"); p.check_token_next(get_comma_tk(), "invalid 'have/assert' declaration, ',' expected");
parser::local_scope scope(p); parser::local_scope scope(p);
is_visible = true;
binder_info bi = mk_contextual_info(is_visible); binder_info bi = mk_contextual_info(is_visible);
expr l = p.save_pos(mk_local(id, prop, bi), pos); expr l = p.save_pos(mk_local(id, prop, bi), pos);
p.add_local(l); p.add_local(l);
@ -478,7 +474,7 @@ static expr parse_have_core(parser & p, pos_info const & pos, optional<expr> con
is_visible = true; is_visible = true;
} else { } else {
p.check_token_next(get_have_tk(), "invalid 'then' declaration, 'have' or 'assert' expected"); p.check_token_next(get_have_tk(), "invalid 'then' declaration, 'have' or 'assert' expected");
is_visible = false; is_visible = true;
} }
body = parse_have_core(p, then_pos, some_expr(l), is_visible); body = parse_have_core(p, then_pos, some_expr(l), is_visible);
} else { } else {

View file

@ -59,7 +59,6 @@ static format * g_let_fmt = nullptr;
static format * g_in_fmt = nullptr; static format * g_in_fmt = nullptr;
static format * g_assign_fmt = nullptr; static format * g_assign_fmt = nullptr;
static format * g_have_fmt = nullptr; static format * g_have_fmt = nullptr;
static format * g_assert_fmt = nullptr;
static format * g_from_fmt = nullptr; static format * g_from_fmt = nullptr;
static format * g_visible_fmt = nullptr; static format * g_visible_fmt = nullptr;
static format * g_show_fmt = nullptr; static format * g_show_fmt = nullptr;
@ -124,7 +123,6 @@ void initialize_pp() {
g_in_fmt = new format(highlight_keyword(format("in"))); g_in_fmt = new format(highlight_keyword(format("in")));
g_assign_fmt = new format(highlight_keyword(format(":="))); g_assign_fmt = new format(highlight_keyword(format(":=")));
g_have_fmt = new format(highlight_keyword(format("have"))); g_have_fmt = new format(highlight_keyword(format("have")));
g_assert_fmt = new format(highlight_keyword(format("assert")));
g_from_fmt = new format(highlight_keyword(format("from"))); g_from_fmt = new format(highlight_keyword(format("from")));
g_visible_fmt = new format(highlight_keyword(format("[visible]"))); g_visible_fmt = new format(highlight_keyword(format("[visible]")));
g_show_fmt = new format(highlight_keyword(format("show"))); g_show_fmt = new format(highlight_keyword(format("show")));
@ -151,7 +149,6 @@ void finalize_pp() {
delete g_in_fmt; delete g_in_fmt;
delete g_assign_fmt; delete g_assign_fmt;
delete g_have_fmt; delete g_have_fmt;
delete g_assert_fmt;
delete g_from_fmt; delete g_from_fmt;
delete g_visible_fmt; delete g_visible_fmt;
delete g_show_fmt; delete g_show_fmt;
@ -779,7 +776,7 @@ auto pretty_fn::pp_have(expr const & e) -> result {
format type_fmt = pp_child(mlocal_type(local), 0).fmt(); format type_fmt = pp_child(mlocal_type(local), 0).fmt();
format proof_fmt = pp_child(proof, 0).fmt(); format proof_fmt = pp_child(proof, 0).fmt();
format body_fmt = pp_child(body, 0).fmt(); format body_fmt = pp_child(body, 0).fmt();
format head_fmt = (binding_info(binding).is_contextual()) ? *g_assert_fmt : *g_have_fmt; format head_fmt = *g_have_fmt;
format r = head_fmt + space() + format(n) + space(); format r = head_fmt + space() + format(n) + space();
r += colon() + nest(m_indent, line() + type_fmt + comma() + space() + *g_from_fmt); r += colon() + nest(m_indent, line() + type_fmt + comma() + space() + *g_from_fmt);
r = group(r); r = group(r);

View file

@ -2,19 +2,18 @@ theorem perm.perm_erase_dup_of_perm [congr] : ∀ {A : Type} [H : decidable_eq A
λ (A : Type) (H : decidable_eq A) (l₁ l₂ : list A) (p : l₁ ~ l₂), λ (A : Type) (H : decidable_eq A) (l₁ l₂ : list A) (p : l₁ ~ l₂),
perm_induction_on p nil perm_induction_on p nil
(λ (x : A) (t₁ t₂ : list A) (p : t₁ ~ t₂) (r : erase_dup t₁ ~ erase_dup t₂), (λ (x : A) (t₁ t₂ : list A) (p : t₁ ~ t₂) (r : erase_dup t₁ ~ erase_dup t₂),
decidable.by_cases decidable.by_cases (λ (xint₁ : x ∈ t₁), have xint₂ : x ∈ t₂, from mem_of_mem_erase_dup …, … …)
(λ (xint₁ : x ∈ t₁), assert xint₂ : x ∈ t₂, from mem_of_mem_erase_dup …, … …)
(λ (nxint₁ : x ∉ t₁), (λ (nxint₁ : x ∉ t₁),
assert nxint₂ : x ∉ t₂, from λ (xint₂ : x ∈ t₂), … nxint₁, have nxint₂ : x ∉ t₂, from λ (xint₂ : x ∈ t₂), … nxint₁,
eq.rec … (eq.symm …))) eq.rec … (eq.symm …)))
(λ (y x : A) (t₁ t₂ : list A) (p : t₁ ~ t₂) (r : erase_dup t₁ ~ erase_dup t₂), (λ (y x : A) (t₁ t₂ : list A) (p : t₁ ~ t₂) (r : erase_dup t₁ ~ erase_dup t₂),
decidable.by_cases decidable.by_cases
(λ (xinyt₁ : x ∈ y :: t₁), (λ (xinyt₁ : x ∈ y :: t₁),
decidable.by_cases (λ (yint₁ : …), …) decidable.by_cases (λ (yint₁ : …), …)
(λ (nyint₁ : y ∉ t₁), assert nyint₂ : …, from …, …)) (λ (nyint₁ : y ∉ t₁), have nyint₂ : …, from …, …))
(λ (nxinyt₁ : x ∉ y :: t₁), (λ (nxinyt₁ : x ∉ y :: t₁),
have xney : x ≠ y, from ne_of_not_mem_cons nxinyt₁, have xney : x ≠ y, from ne_of_not_mem_cons nxinyt₁,
have nxint₁ : x ∉ t₁, from not_mem_of_not_mem_cons nxinyt₁, have nxint₁ : x ∉ t₁, from not_mem_of_not_mem_cons nxinyt₁,
assert nxint₂ : x ∉ t₂, from λ (xint₂ : …), …, have nxint₂ : x ∉ t₂, from λ (xint₂ : …), …,
… …)) … …))
(λ (t₁ t₂ t₃ : list A) (p₁ : t₁ ~ t₂) (p₂ : t₂ ~ t₃), trans) (λ (t₁ t₂ t₃ : list A) (p₁ : t₁ ~ t₂) (p₂ : t₂ ~ t₃), trans)

View file

@ -9,6 +9,7 @@ a : A,
l₁ l₂ : list A, l₁ l₂ : list A,
h : subcount (a :: l₁) l₂ = tt, h : subcount (a :: l₁) l₂ = tt,
x : A, x : A,
this : subcount l₁ l₂ = tt,
ih : ∀ (a : A), list.count a l₁ ≤ list.count a l₂, ih : ∀ (a : A), list.count a l₁ ≤ list.count a l₂,
i : list.count a (a :: l₁) ≤ list.count a l₂, i : list.count a (a :: l₁) ≤ list.count a l₂,
this : x = a this : x = a

View file

@ -65,9 +65,9 @@ remark: set 'formatter.hide_full_terms' to false to see the complete term
(λ (l : …), (λ (l : …),
@list.induction_on A … l … @list.induction_on A … l …
(λ (a : A) (l' : …) (IH : …) (nodup_al' : …), (λ (a : A) (l' : …) (IH : …) (nodup_al' : …),
assert anl' : …, from …, have anl' : …, from …,
assert H3 : …, from …, have H3 : …, from …,
assert nodup_l' : …, from …, have nodup_l' : …, from …,
assert P_l' : …, from …, have P_l' : …, from …,
?M_1 …))) ?M_1 …)))
finset_induction_bug.lean:30:49: error: invalid end of module, expecting 'end' finset_induction_bug.lean:30:49: error: invalid end of module, expecting 'end'

View file

@ -5,10 +5,10 @@ constants a b c : bool
axiom H1 : a = b axiom H1 : a = b
axiom H2 : b = c axiom H2 : b = c
check have e1 [visible] : a = b, from H1, check have e1 : a = b, from H1,
have e2 : a = c, by apply trans; apply e1; apply H2, have e2 : a = c, by apply trans; apply e1; apply H2,
have e3 : c = a, from e2⁻¹, have e3 : c = a, from e2⁻¹,
have e4 [visible] : b = a, from e1⁻¹, have e4 : b = a, from e1⁻¹,
have e5 : b = c, from e4 ⬝ e2, have e5 : b = c, from e4 ⬝ e2,
have e6 : a = a, from H1 ⬝ H2 ⬝ H2⁻¹ ⬝ H1⁻¹ ⬝ H1 ⬝ H2 ⬝ H2⁻¹ ⬝ H1⁻¹, have e6 : a = a, from H1 ⬝ H2 ⬝ H2⁻¹ ⬝ H1⁻¹ ⬝ H1 ⬝ H2 ⬝ H2⁻¹ ⬝ H1⁻¹,
e3 ⬝ e2 e3 ⬝ e2

View file

@ -1,7 +1,7 @@
assert e1 : a = b, from H1, have e1 : a = b, from H1,
have e2 : a = c, from e1 ⬝ H2, have e2 : a = c, from e1 ⬝ H2,
have e3 : c = a, from e2⁻¹, have e3 : c = a, from e2⁻¹,
assert e4 : b = a, from e1⁻¹, have e4 : b = a, from e1⁻¹,
have e5 : b = c, from e4 ⬝ e2, have e5 : b = c, from e4 ⬝ e2,
have e6 : a = a, from H1 ⬝ H2 ⬝ H2⁻¹ ⬝ H1⁻¹ ⬝ H1 ⬝ H2 ⬝ H2⁻¹ ⬝ H1⁻¹, have e6 : a = a, from H1 ⬝ H2 ⬝ H2⁻¹ ⬝ H1⁻¹ ⬝ H1 ⬝ H2 ⬝ H2⁻¹ ⬝ H1⁻¹,
e3 ⬝ e2 : e3 ⬝ e2 :

View file

@ -23,7 +23,7 @@ show P, begin+ exact H end
example : P := example : P :=
have H : P, from HP, have H : P, from HP,
show P, using H, by exact H show P, by exact H
example : P := example : P :=
assert H : P, from HP, assert H : P, from HP,

View file

@ -5,5 +5,5 @@ axiom Ha : a
axiom Hb : b axiom Hb : b
axiom Hc : c axiom Hc : c
check have H1 : a, from Ha, check have H1 : a, from Ha,
have H2 : a, using H1, from H1, have H2 : a, from H1,
H2 H2

View file

@ -3,7 +3,7 @@ assume Hpqr : p ∧ q ∧ r,
have Hp : p, from and.elim_left Hpqr, have Hp : p, from and.elim_left Hpqr,
have Hqr : q ∧ r, from and.elim_right Hpqr, have Hqr : q ∧ r, from and.elim_right Hpqr,
have Hq : q, from and.elim_left Hqr, have Hq : q, from and.elim_left Hqr,
show q ∧ p, using Hp Hq, from show q ∧ p, from
proof proof
and.intro Hq Hp and.intro Hq Hp
qed qed

View file

@ -3,10 +3,11 @@ variable {f : A → A → A}
variable {finv : A → A} variable {finv : A → A}
premise (h : ∀ x y : A, finv (f x y) = y) premise (h : ∀ x y : A, finv (f x y) = y)
include h
theorem foo₁ : ∀ x y z : A, f x y = f x z → y = z := theorem foo₁ : ∀ x y z : A, f x y = f x z → y = z :=
λ x y z, assume e, using h, from sorry λ x y z, assume e, sorry
theorem foo₂ : ∀ x y z : A, f x y = f x z → y = z := theorem foo₂ : ∀ x y z : A, f x y = f x z → y = z :=
λ x y z, assume e, λ x y z, assume e,
assert s₁ : finv (f x y) = finv (f x z), by rewrite e, assert s₁ : finv (f x y) = finv (f x z), by rewrite e,
show y = z, using h, by rewrite *h at s₁; exact s₁ show y = z, by rewrite *h at s₁; exact s₁

View file

@ -1,8 +1,8 @@
show a = c, from H1 ⬝ H2 : a = c show a = c, from H1 ⬝ H2 : a = c
------------ ------------
assert e1 : a = b, from H1, have e1 : a = b, from H1,
have e2 : a = c, from e1 ⬝ H2, have e2 : a = c, from e1 ⬝ H2,
have e3 : c = a, from e2⁻¹, have e3 : c = a, from e2⁻¹,
assert e4 : b = a, from e1⁻¹, have e4 : b = a, from e1⁻¹,
show b = c, from e1⁻¹ ⬝ e2 : show b = c, from e1⁻¹ ⬝ e2 :
b = c b = c