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,19 +214,11 @@ 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 sup_fn_interval_aux1 : ∀ ε : , ε > 0 → f (inter_sup a b f) ≥ 0 - ε :=
(take ε, suppose ε > 0,
have ¬ f (inter_sup a b f) < 0 - ε, from
(suppose f (inter_sup a b f) < 0 - ε,
private theorem sup_fn_interval_aux1 : f (inter_sup a b f) ≥ 0 :=
have ¬ f (inter_sup a b f) < 0, from
(suppose f (inter_sup a b f) < 0,
have ∃ δ, δ > (0 : ) ∧ inter_sup a b f + δ < b ∧ f (inter_sup a b f + δ) < 0, begin
let Hcont := Hf (inter_sup a b f)
(show -f (inter_sup a b f) > 0, begin
rewrite zero_sub at this,
apply lt.trans,
exact `ε > 0`,
apply lt_neg_of_lt_neg,
exact this
end),
let Hcont := Hf (inter_sup a b f) (neg_pos_of_neg this),
cases Hcont with δ Hδ,
cases em (inter_sup a b f + δ / 2 < b),
existsi δ / 2,
@ -270,14 +262,13 @@ private theorem sup_fn_interval_aux1 : ∀ ε : , ε > 0 → f (inter_sup a b
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)
le_of_not_gt this
private theorem sup_fn_interval_aux2 : ∀ ε : , ε > 0 → f (inter_sup a b f) < ε :=
(take ε, suppose ε > 0,
have ¬ f (inter_sup a b f) ≥ ε, from
(suppose f (inter_sup a b f) ≥ ε,
private theorem sup_fn_interval_aux2 : f (inter_sup a b f) ≤ 0 :=
have ¬ f (inter_sup a b f) > 0, from
(suppose f (inter_sup a b f) > 0,
have ∃ δ : , δ > 0 ∧ ∀ x' : , abs (inter_sup a b f - x') < δ → f x' > 0, begin
let Hcont := Hf (inter_sup a b f) (show f (inter_sup a b f) > 0, from gt_of_ge_of_gt this `ε > 0`),
let Hcont := Hf (inter_sup a b f) this,
cases Hcont with δ Hδ,
existsi δ,
split,
@ -314,12 +305,12 @@ private theorem sup_fn_interval_aux2 : ∀ ε : , ε > 0 → f (inter_sup a b
show false, from (not_lt_of_gt this) (and.right `x ∈ {x | x < b ∧ f x < 0}`)),
le_of_not_gt 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 (sub_lt_of_pos _ (div_pos_of_pos_of_pos (and.left Hδ) (two_pos))) Hle),
lt_of_not_ge this)
show false, from not_le_of_gt
(sub_lt_of_pos _ (div_pos_of_pos_of_pos (and.left Hδ) (two_pos))) Hle),
le_of_not_gt this
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_zero_of_nonneg_of_forall_lt Hnlt sup_fn_interval_aux2
eq_of_le_of_ge sup_fn_interval_aux2 sup_fn_interval_aux1
private theorem intermediate_value_incr_aux1 : ∃ δ : , δ > 0 ∧ ∀ y, abs (b - y) < δ → f y > 0 :=
begin
@ -448,6 +439,7 @@ theorem intermediate_value_decr {f : } (Hf : continuous f) {a b :
end
end inter_val
/-
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₂ :=