Theory HOL-Library.Extended_Real
section ‹Extended real number line›
theory Extended_Real
imports Complex_Main Extended_Nat Liminf_Limsup
begin
text ‹
This should be part of \<^theory>‹HOL-Library.Extended_Nat› or \<^theory>‹HOL-Library.Order_Continuity›, but then the AFP-entry ‹Jinja_Thread› fails, as it does overload
certain named from \<^theory>‹Complex_Main›.
›
lemma incseq_sumI2:
fixes f :: "'i ⇒ nat ⇒ 'a::ordered_comm_monoid_add"
shows "(⋀n. n ∈ A ⟹ mono (f n)) ⟹ mono (λi. ∑n∈A. f n i)"
unfolding incseq_def by (auto intro: sum_mono)
lemma incseq_sumI:
fixes f :: "nat ⇒ 'a::ordered_comm_monoid_add"
assumes "⋀i. 0 ≤ f i"
shows "incseq (λi. sum f {..< i})"
proof (intro incseq_SucI)
fix n
have "sum f {..< n} + 0 ≤ sum f {..<n} + f n"
using assms by (rule add_left_mono)
then show "sum f {..< n} ≤ sum f {..< Suc n}"
by auto
qed
lemma continuous_at_left_imp_sup_continuous:
fixes f :: "'a::{complete_linorder, linorder_topology} ⇒ 'b::{complete_linorder, linorder_topology}"
assumes "mono f" "⋀x. continuous (at_left x) f"
shows "sup_continuous f"
unfolding sup_continuous_def
proof safe
fix M :: "nat ⇒ 'a" assume "incseq M" then show "f (SUP i. M i) = (SUP i. f (M i))"
using continuous_at_Sup_mono [OF assms, of "range M"] by (simp add: image_comp)
qed
lemma sup_continuous_at_left:
fixes f :: "'a::{complete_linorder, linorder_topology, first_countable_topology} ⇒
'b::{complete_linorder, linorder_topology}"
assumes f: "sup_continuous f"
shows "continuous (at_left x) f"
proof cases
assume "x = bot" then show ?thesis
by (simp add: trivial_limit_at_left_bot)
next
assume x: "x ≠ bot"
show ?thesis
unfolding continuous_within
proof (intro tendsto_at_left_sequentially[of bot])
fix S :: "nat ⇒ 'a" assume S: "incseq S" and S_x: "S ⇢ x"
from S_x have x_eq: "x = (SUP i. S i)"
by (rule LIMSEQ_unique) (intro LIMSEQ_SUP S)
show "(λn. f (S n)) ⇢ f x"
unfolding x_eq sup_continuousD[OF f S]
using S sup_continuous_mono[OF f] by (intro LIMSEQ_SUP) (auto simp: mono_def)
qed (insert x, auto simp: bot_less)
qed
lemma sup_continuous_iff_at_left:
fixes f :: "'a::{complete_linorder, linorder_topology, first_countable_topology} ⇒
'b::{complete_linorder, linorder_topology}"
shows "sup_continuous f ⟷ (∀x. continuous (at_left x) f) ∧ mono f"
using sup_continuous_at_left[of f] continuous_at_left_imp_sup_continuous[of f]
sup_continuous_mono[of f] by auto
lemma continuous_at_right_imp_inf_continuous:
fixes f :: "'a::{complete_linorder, linorder_topology} ⇒ 'b::{complete_linorder, linorder_topology}"
assumes "mono f" "⋀x. continuous (at_right x) f"
shows "inf_continuous f"
unfolding inf_continuous_def
proof safe
fix M :: "nat ⇒ 'a" assume "decseq M" then show "f (INF i. M i) = (INF i. f (M i))"
using continuous_at_Inf_mono [OF assms, of "range M"] by (simp add: image_comp)
qed
lemma inf_continuous_at_right:
fixes f :: "'a::{complete_linorder, linorder_topology, first_countable_topology} ⇒
'b::{complete_linorder, linorder_topology}"
assumes f: "inf_continuous f"
shows "continuous (at_right x) f"
proof cases
assume "x = top" then show ?thesis
by (simp add: trivial_limit_at_right_top)
next
assume x: "x ≠ top"
show ?thesis
unfolding continuous_within
proof (intro tendsto_at_right_sequentially[of _ top])
fix S :: "nat ⇒ 'a" assume S: "decseq S" and S_x: "S ⇢ x"
from S_x have x_eq: "x = (INF i. S i)"
by (rule LIMSEQ_unique) (intro LIMSEQ_INF S)
show "(λn. f (S n)) ⇢ f x"
unfolding x_eq inf_continuousD[OF f S]
using S inf_continuous_mono[OF f] by (intro LIMSEQ_INF) (auto simp: mono_def antimono_def)
qed (insert x, auto simp: less_top)
qed
lemma inf_continuous_iff_at_right:
fixes f :: "'a::{complete_linorder, linorder_topology, first_countable_topology} ⇒
'b::{complete_linorder, linorder_topology}"
shows "inf_continuous f ⟷ (∀x. continuous (at_right x) f) ∧ mono f"
using inf_continuous_at_right[of f] continuous_at_right_imp_inf_continuous[of f]
inf_continuous_mono[of f] by auto
instantiation enat :: linorder_topology
begin
definition open_enat :: "enat set ⇒ bool" where
"open_enat = generate_topology (range lessThan ∪ range greaterThan)"
instance
proof qed (rule open_enat_def)
end
lemma open_enat: "open {enat n}"
proof (cases n)
case 0
then have "{enat n} = {..< eSuc 0}"
by (auto simp: enat_0)
then show ?thesis
by simp
next
case (Suc n')
then have "{enat n} = {enat n' <..< enat (Suc n)}"
using enat_iless by (fastforce simp: set_eq_iff)
then show ?thesis
by simp
qed
lemma open_enat_iff:
fixes A :: "enat set"
shows "open A ⟷ (∞ ∈ A ⟶ (∃n::nat. {n <..} ⊆ A))"
proof safe
assume "∞ ∉ A"
then have "A = (⋃n∈{n. enat n ∈ A}. {enat n})"
by (simp add: set_eq_iff) (metis not_enat_eq)
moreover have "open …"
by (auto intro: open_enat)
ultimately show "open A"
by simp
next
fix n assume "{enat n <..} ⊆ A"
then have "A = (⋃n∈{n. enat n ∈ A}. {enat n}) ∪ {enat n <..}"
using enat_ile leI by (simp add: set_eq_iff) blast
moreover have "open …"
by (intro open_Un open_UN ballI open_enat open_greaterThan)
ultimately show "open A"
by simp
next
assume "open A" "∞ ∈ A"
then have "generate_topology (range lessThan ∪ range greaterThan) A" "∞ ∈ A"
unfolding open_enat_def by auto
then show "∃n::nat. {n <..} ⊆ A"
proof induction
case (Int A B)
then obtain n m where "{enat n<..} ⊆ A" "{enat m<..} ⊆ B"
by auto
then have "{enat (max n m) <..} ⊆ A ∩ B"
by (auto simp add: subset_eq Ball_def max_def simp flip: enat_ord_code(1))
then show ?case
by auto
next
case (UN K)
then obtain k where "k ∈ K" "∞ ∈ k"
by auto
with UN.IH[OF this] show ?case
by auto
qed auto
qed
lemma nhds_enat: "nhds x = (if x = ∞ then INF i. principal {enat i..} else principal {x})"
proof auto
show "nhds ∞ = (INF i. principal {enat i..})"
proof (rule antisym)
show "nhds ∞ ≤ (INF i. principal {enat i..})"
unfolding nhds_def
using Ioi_le_Ico by (intro INF_greatest INF_lower) (auto simp add: open_enat_iff)
show "(INF i. principal {enat i..}) ≤ nhds ∞"
unfolding nhds_def
by (intro INF_greatest) (force intro: INF_lower2[of "Suc _"] simp add: open_enat_iff Suc_ile_eq)
qed
show "nhds (enat i) = principal {enat i}" for i
by (simp add: nhds_discrete_open open_enat)
qed
instance enat :: topological_comm_monoid_add
proof
have [simp]: "enat i ≤ aa ⟹ enat i ≤ aa + ba" for aa ba i
by (rule order_trans[OF _ add_mono[of aa aa 0 ba]]) auto
then have [simp]: "enat i ≤ ba ⟹ enat i ≤ aa + ba" for aa ba i
by (metis add.commute)
fix a b :: enat show "((λx. fst x + snd x) ⤏ a + b) (nhds a ×⇩F nhds b)"
apply (auto simp: nhds_enat filterlim_INF prod_filter_INF1 prod_filter_INF2
filterlim_principal principal_prod_principal eventually_principal)
subgoal for i
by (auto intro!: eventually_INF1[of i] simp: eventually_principal)
subgoal for j i
by (auto intro!: eventually_INF1[of i] simp: eventually_principal)
subgoal for j i
by (auto intro!: eventually_INF1[of i] simp: eventually_principal)
done
qed
text ‹
For more lemmas about the extended real numbers see
🗏‹~~/src/HOL/Analysis/Extended_Real_Limits.thy›.
›
subsection ‹Definition and basic properties›