feat(library/theories/analysis): clean and simplify proof of IVT

This commit is contained in:
Rob Lewis 2015-09-15 21:06:11 -04:00 committed by Leonardo de Moura
parent ee257a7c6c
commit 631b9b3312

View file

@ -214,112 +214,103 @@ private theorem Hmem : ∀ x, x ∈ {x | x < b ∧ f x < 0} → x ≤ b := λ x
private theorem Hsupleb : inter_sup a b f ≤ b := sup_le (Hinh) Hmem private theorem Hsupleb : inter_sup a b f ≤ b := sup_le (Hinh) Hmem
private theorem sup_fn_interval_aux1 : ∀ ε : , ε > 0 → f (inter_sup a b f) ≥ 0 - ε := private theorem sup_fn_interval_aux1 : f (inter_sup a b f) ≥ 0 :=
(take ε, suppose ε > 0, have ¬ f (inter_sup a b f) < 0, from
have ¬ f (inter_sup a b f) < 0 - ε, from (suppose f (inter_sup a b f) < 0,
(suppose f (inter_sup a b f) < 0 - ε, have ∃ δ, δ > (0 : ) ∧ inter_sup a b f + δ < b ∧ f (inter_sup a b f + δ) < 0, begin
have ∃ δ, δ > (0 : ) ∧ inter_sup a b f + δ < b ∧ f (inter_sup a b f + δ) < 0, begin let Hcont := Hf (inter_sup a b f) (neg_pos_of_neg this),
let Hcont := Hf (inter_sup a b f) cases Hcont with δ Hδ,
(show -f (inter_sup a b f) > 0, begin cases em (inter_sup a b f + δ / 2 < b),
rewrite zero_sub at this, existsi δ / 2,
apply lt.trans, split,
exact `ε > 0`, apply div_pos_of_pos_of_pos (and.left Hδ) two_pos,
apply lt_neg_of_lt_neg, split,
exact this assumption,
end), have Habs : abs (inter_sup a b f - (inter_sup a b f + δ / 2)) < δ, begin
cases Hcont with δ Hδ, krewrite [sub_add_eq_sub_sub, sub_self, zero_sub, abs_neg,
cases em (inter_sup a b f + δ / 2 < b), abs_of_pos (div_pos_of_pos_of_pos (and.left Hδ) two_pos)],
existsi δ / 2, apply div_two_lt_of_pos (and.left Hδ)
split, end,
apply div_pos_of_pos_of_pos (and.left Hδ) two_pos, let Hlt := and.right Hδ _ Habs,
split, let Hlt' := sub_lt_of_abs_sub_lt_left Hlt,
assumption, let Hlt'' := lt_add_of_sub_lt_right _ _ _ Hlt',
have Habs : abs (inter_sup a b f - (inter_sup a b f + δ / 2)) < δ, begin rewrite [-sub_eq_add_neg at Hlt'', sub_self at Hlt''],
krewrite [sub_add_eq_sub_sub, sub_self, zero_sub, abs_neg, assumption,
abs_of_pos (div_pos_of_pos_of_pos (and.left Hδ) two_pos)], let Hble := le_of_not_gt a_1,
apply div_two_lt_of_pos (and.left Hδ) have Habs : abs (inter_sup a b f - b) < δ, begin
end, apply abs_lt_of_lt_of_neg_lt,
let Hlt := and.right Hδ _ Habs, apply sub_lt_left_of_lt_add,
let Hlt' := sub_lt_of_abs_sub_lt_left Hlt, apply lt_of_le_of_lt,
let Hlt'' := lt_add_of_sub_lt_right _ _ _ Hlt', apply Hsupleb,
rewrite [-sub_eq_add_neg at Hlt'', sub_self at Hlt''], apply lt_add_of_pos_right (and.left Hδ),
assumption, rewrite neg_sub,
let Hble := le_of_not_gt a_1, apply sub_lt_left_of_lt_add,
have Habs : abs (inter_sup a b f - b) < δ, begin apply lt_of_le_of_lt,
apply abs_lt_of_lt_of_neg_lt, apply Hble,
apply sub_lt_left_of_lt_add, apply add_lt_add_left,
apply lt_of_le_of_lt, apply div_two_lt_of_pos (and.left Hδ)
apply Hsupleb, end,
apply lt_add_of_pos_right (and.left Hδ), let Hlt := and.right Hδ _ Habs,
rewrite neg_sub, let Hlt' := sub_lt_of_abs_sub_lt_left Hlt,
apply sub_lt_left_of_lt_add, let Hlt'' := lt_add_of_sub_lt_right _ _ _ Hlt',
apply lt_of_le_of_lt, rewrite [-sub_eq_add_neg at Hlt'', sub_self at Hlt''],
apply Hble, apply absurd Hb (not_lt_of_ge (le_of_lt Hlt''))
apply add_lt_add_left, end,
apply div_two_lt_of_pos (and.left Hδ) obtain δ Hδ, from this,
end, have inter_sup a b f + δ ∈ {x | x < b ∧ f x < 0},
let Hlt := and.right Hδ _ Habs, from and.intro (and.left (and.right Hδ)) (and.right (and.right Hδ)),
let Hlt' := sub_lt_of_abs_sub_lt_left Hlt, have Hle : ¬ inter_sup a b f < inter_sup a b f + δ,
let Hlt'' := lt_add_of_sub_lt_right _ _ _ Hlt', from not_lt_of_ge (le_sup this Hmem),
rewrite [-sub_eq_add_neg at Hlt'', sub_self at Hlt''], show false, from Hle (lt_add_of_pos_right (and.left Hδ))),
apply absurd Hb (not_lt_of_ge (le_of_lt Hlt'')) le_of_not_gt this
end,
obtain δ Hδ, from this,
have inter_sup a b f + δ ∈ {x | x < b ∧ f x < 0},
from and.intro (and.left (and.right Hδ)) (and.right (and.right Hδ)),
have Hle : ¬ inter_sup a b f < inter_sup a b f + δ,
from not_lt_of_ge (le_sup this Hmem),
show false, from Hle (lt_add_of_pos_right (and.left Hδ))),
le_of_not_gt this)
private theorem sup_fn_interval_aux2 : ∀ ε : , ε > 0 → f (inter_sup a b f) < ε := private theorem sup_fn_interval_aux2 : f (inter_sup a b f) ≤ 0 :=
(take ε, suppose ε > 0, have ¬ f (inter_sup a b f) > 0, from
have ¬ f (inter_sup a b f) ≥ ε, from (suppose f (inter_sup a b f) > 0,
(suppose f (inter_sup a b f) ≥ ε, have ∃ δ : , δ > 0 ∧ ∀ x' : , abs (inter_sup a b f - x') < δ → f x' > 0, begin
have ∃ δ : , δ > 0 ∧ ∀ x' : , abs (inter_sup a b f - x') < δ → f x' > 0, begin let Hcont := Hf (inter_sup a b f) this,
let Hcont := Hf (inter_sup a b f) (show f (inter_sup a b f) > 0, from gt_of_ge_of_gt this `ε > 0`), cases Hcont with δ Hδ,
cases Hcont with δ Hδ, existsi δ,
existsi δ, split,
split, exact and.left Hδ,
exact and.left Hδ, intro x Hx,
intro x Hx, let Hlt := and.right Hδ _ Hx,
let Hlt := and.right Hδ _ Hx, let Hlt' := sub_lt_of_abs_sub_lt_right Hlt,
let Hlt' := sub_lt_of_abs_sub_lt_right Hlt, rewrite sub_self at Hlt',
rewrite sub_self at Hlt', exact Hlt'
exact Hlt' end,
end, obtain δ Hδ, from this,
obtain δ Hδ, from this, have ∀ x, x ∈ {x | x < b ∧ f x < 0} → x ≤ inter_sup a b f - δ / 2, from
have ∀ x, x ∈ {x | x < b ∧ f x < 0} → x ≤ inter_sup a b f - δ / 2, from (take x, suppose x ∈ {x | x < b ∧ f x < 0},
(take x, suppose x ∈ {x | x < b ∧ f x < 0}, have x ≤ inter_sup a b f, from le_sup this Hmem,
have x ≤ inter_sup a b f, from le_sup this Hmem, have ¬ x > inter_sup a b f - δ / 2, from
have ¬ x > inter_sup a b f - δ / 2, from (assume Hngt,
(assume Hngt, have abs (inter_sup a b f - x) < δ, begin
have abs (inter_sup a b f - x) < δ, begin apply abs_lt_of_lt_of_neg_lt,
apply abs_lt_of_lt_of_neg_lt, apply sub_lt_of_sub_lt,
apply sub_lt_of_sub_lt, apply gt.trans,
apply gt.trans, exact Hngt,
exact Hngt, apply sub_lt_sub_left,
apply sub_lt_sub_left, exact div_two_lt_of_pos (and.left Hδ),
exact div_two_lt_of_pos (and.left Hδ), rewrite neg_sub,
rewrite neg_sub, apply lt_of_le_of_lt,
apply lt_of_le_of_lt, rotate 1,
rotate 1, apply and.left Hδ,
apply and.left Hδ, apply sub_nonpos_of_le,
apply sub_nonpos_of_le, apply le_sup,
apply le_sup, exact this,
exact this, exact Hmem
exact Hmem end,
end, have f x > 0, from and.right Hδ _ this,
have f x > 0, from and.right Hδ _ this, show false, from (not_lt_of_gt this) (and.right `x ∈ {x | x < b ∧ f x < 0}`)),
show false, from (not_lt_of_gt this) (and.right `x ∈ {x | x < b ∧ f x < 0}`)), le_of_not_gt this),
le_of_not_gt this), have Hle : inter_sup a b f ≤ inter_sup a b f - δ / 2, from sup_le Hinh this,
have Hle : inter_sup a b f ≤ inter_sup a b f - δ / 2, from sup_le Hinh this, show false, from not_le_of_gt
show false, from not_le_of_gt (sub_lt_of_pos _ (div_pos_of_pos_of_pos (and.left Hδ) (two_pos))) Hle), (sub_lt_of_pos _ (div_pos_of_pos_of_pos (and.left Hδ) (two_pos))) Hle),
lt_of_not_ge this) le_of_not_gt this
private theorem sup_fn_interval : f (inter_sup a b f) = 0 := private theorem sup_fn_interval : f (inter_sup a b f) = 0 :=
have Hnlt : f (inter_sup a b f) ≥ 0, from ge_of_forall_ge_sub sup_fn_interval_aux1, eq_of_le_of_ge sup_fn_interval_aux2 sup_fn_interval_aux1
eq_zero_of_nonneg_of_forall_lt Hnlt sup_fn_interval_aux2
private theorem intermediate_value_incr_aux1 : ∃ δ : , δ > 0 ∧ ∀ y, abs (b - y) < δ → f y > 0 := private theorem intermediate_value_incr_aux1 : ∃ δ : , δ > 0 ∧ ∀ y, abs (b - y) < δ → f y > 0 :=
begin begin
@ -448,6 +439,7 @@ theorem intermediate_value_decr {f : } (Hf : continuous f) {a b :
end end
end inter_val end inter_val
/- /-
proposition converges_to_at_unique {f : M → N} {y₁ y₂ : N} {x : M} proposition converges_to_at_unique {f : M → N} {y₁ y₂ : N} {x : M}
(H₁ : f ⟶ y₁ '[at x]) (H₂ : f ⟶ y₂ '[at x]) : y₁ = y₂ := (H₁ : f ⟶ y₁ '[at x]) (H₂ : f ⟶ y₂ '[at x]) : y₁ = y₂ :=