diff --git a/src/HOL/Filter.thy b/src/HOL/Filter.thy --- a/src/HOL/Filter.thy +++ b/src/HOL/Filter.thy @@ -1,1996 +1,1998 @@ (* Title: HOL/Filter.thy Author: Brian Huffman Author: Johannes Hölzl *) section \Filters on predicates\ theory Filter imports Set_Interval Lifting_Set begin subsection \Filters\ text \ This definition also allows non-proper filters. \ locale is_filter = fixes F :: "('a \ bool) \ bool" assumes True: "F (\x. True)" assumes conj: "F (\x. P x) \ F (\x. Q x) \ F (\x. P x \ Q x)" assumes mono: "\x. P x \ Q x \ F (\x. P x) \ F (\x. Q x)" typedef 'a filter = "{F :: ('a \ bool) \ bool. is_filter F}" proof show "(\x. True) \ ?filter" by (auto intro: is_filter.intro) qed lemma is_filter_Rep_filter: "is_filter (Rep_filter F)" using Rep_filter [of F] by simp lemma Abs_filter_inverse': assumes "is_filter F" shows "Rep_filter (Abs_filter F) = F" using assms by (simp add: Abs_filter_inverse) subsubsection \Eventually\ definition eventually :: "('a \ bool) \ 'a filter \ bool" where "eventually P F \ Rep_filter F P" syntax "_eventually" :: "pttrn => 'a filter => bool => bool" ("(3\\<^sub>F _ in _./ _)" [0, 0, 10] 10) translations "\\<^sub>Fx in F. P" == "CONST eventually (\x. P) F" lemma eventually_Abs_filter: assumes "is_filter F" shows "eventually P (Abs_filter F) = F P" unfolding eventually_def using assms by (simp add: Abs_filter_inverse) lemma filter_eq_iff: shows "F = F' \ (\P. eventually P F = eventually P F')" unfolding Rep_filter_inject [symmetric] fun_eq_iff eventually_def .. lemma eventually_True [simp]: "eventually (\x. True) F" unfolding eventually_def by (rule is_filter.True [OF is_filter_Rep_filter]) lemma always_eventually: "\x. P x \ eventually P F" proof - assume "\x. P x" hence "P = (\x. True)" by (simp add: ext) thus "eventually P F" by simp qed lemma eventuallyI: "(\x. P x) \ eventually P F" by (auto intro: always_eventually) lemma eventually_mono: "\eventually P F; \x. P x \ Q x\ \ eventually Q F" unfolding eventually_def by (blast intro: is_filter.mono [OF is_filter_Rep_filter]) lemma eventually_conj: assumes P: "eventually (\x. P x) F" assumes Q: "eventually (\x. Q x) F" shows "eventually (\x. P x \ Q x) F" using assms unfolding eventually_def by (rule is_filter.conj [OF is_filter_Rep_filter]) lemma eventually_mp: assumes "eventually (\x. P x \ Q x) F" assumes "eventually (\x. P x) F" shows "eventually (\x. Q x) F" proof - have "eventually (\x. (P x \ Q x) \ P x) F" using assms by (rule eventually_conj) then show ?thesis by (blast intro: eventually_mono) qed lemma eventually_rev_mp: assumes "eventually (\x. P x) F" assumes "eventually (\x. P x \ Q x) F" shows "eventually (\x. Q x) F" using assms(2) assms(1) by (rule eventually_mp) lemma eventually_conj_iff: "eventually (\x. P x \ Q x) F \ eventually P F \ eventually Q F" by (auto intro: eventually_conj elim: eventually_rev_mp) lemma eventually_elim2: assumes "eventually (\i. P i) F" assumes "eventually (\i. Q i) F" assumes "\i. P i \ Q i \ R i" shows "eventually (\i. R i) F" using assms by (auto elim!: eventually_rev_mp) lemma eventually_ball_finite_distrib: "finite A \ (eventually (\x. \y\A. P x y) net) \ (\y\A. eventually (\x. P x y) net)" by (induction A rule: finite_induct) (auto simp: eventually_conj_iff) lemma eventually_ball_finite: "finite A \ \y\A. eventually (\x. P x y) net \ eventually (\x. \y\A. P x y) net" by (auto simp: eventually_ball_finite_distrib) lemma eventually_all_finite: fixes P :: "'a \ 'b::finite \ bool" assumes "\y. eventually (\x. P x y) net" shows "eventually (\x. \y. P x y) net" using eventually_ball_finite [of UNIV P] assms by simp lemma eventually_ex: "(\\<^sub>Fx in F. \y. P x y) \ (\Y. \\<^sub>Fx in F. P x (Y x))" proof assume "\\<^sub>Fx in F. \y. P x y" then have "\\<^sub>Fx in F. P x (SOME y. P x y)" by (auto intro: someI_ex eventually_mono) then show "\Y. \\<^sub>Fx in F. P x (Y x)" by auto qed (auto intro: eventually_mono) lemma not_eventually_impI: "eventually P F \ \ eventually Q F \ \ eventually (\x. P x \ Q x) F" by (auto intro: eventually_mp) lemma not_eventuallyD: "\ eventually P F \ \x. \ P x" by (metis always_eventually) lemma eventually_subst: assumes "eventually (\n. P n = Q n) F" shows "eventually P F = eventually Q F" (is "?L = ?R") proof - from assms have "eventually (\x. P x \ Q x) F" and "eventually (\x. Q x \ P x) F" by (auto elim: eventually_mono) then show ?thesis by (auto elim: eventually_elim2) qed subsection \ Frequently as dual to eventually \ definition frequently :: "('a \ bool) \ 'a filter \ bool" where "frequently P F \ \ eventually (\x. \ P x) F" syntax "_frequently" :: "pttrn \ 'a filter \ bool \ bool" ("(3\\<^sub>F _ in _./ _)" [0, 0, 10] 10) translations "\\<^sub>Fx in F. P" == "CONST frequently (\x. P) F" lemma not_frequently_False [simp]: "\ (\\<^sub>Fx in F. False)" by (simp add: frequently_def) lemma frequently_ex: "\\<^sub>Fx in F. P x \ \x. P x" by (auto simp: frequently_def dest: not_eventuallyD) lemma frequentlyE: assumes "frequently P F" obtains x where "P x" using frequently_ex[OF assms] by auto lemma frequently_mp: assumes ev: "\\<^sub>Fx in F. P x \ Q x" and P: "\\<^sub>Fx in F. P x" shows "\\<^sub>Fx in F. Q x" proof - from ev have "eventually (\x. \ Q x \ \ P x) F" by (rule eventually_rev_mp) (auto intro!: always_eventually) from eventually_mp[OF this] P show ?thesis by (auto simp: frequently_def) qed lemma frequently_rev_mp: assumes "\\<^sub>Fx in F. P x" assumes "\\<^sub>Fx in F. P x \ Q x" shows "\\<^sub>Fx in F. Q x" using assms(2) assms(1) by (rule frequently_mp) lemma frequently_mono: "(\x. P x \ Q x) \ frequently P F \ frequently Q F" using frequently_mp[of P Q] by (simp add: always_eventually) lemma frequently_elim1: "\\<^sub>Fx in F. P x \ (\i. P i \ Q i) \ \\<^sub>Fx in F. Q x" by (metis frequently_mono) lemma frequently_disj_iff: "(\\<^sub>Fx in F. P x \ Q x) \ (\\<^sub>Fx in F. P x) \ (\\<^sub>Fx in F. Q x)" by (simp add: frequently_def eventually_conj_iff) lemma frequently_disj: "\\<^sub>Fx in F. P x \ \\<^sub>Fx in F. Q x \ \\<^sub>Fx in F. P x \ Q x" by (simp add: frequently_disj_iff) lemma frequently_bex_finite_distrib: assumes "finite A" shows "(\\<^sub>Fx in F. \y\A. P x y) \ (\y\A. \\<^sub>Fx in F. P x y)" using assms by induction (auto simp: frequently_disj_iff) lemma frequently_bex_finite: "finite A \ \\<^sub>Fx in F. \y\A. P x y \ \y\A. \\<^sub>Fx in F. P x y" by (simp add: frequently_bex_finite_distrib) lemma frequently_all: "(\\<^sub>Fx in F. \y. P x y) \ (\Y. \\<^sub>Fx in F. P x (Y x))" using eventually_ex[of "\x y. \ P x y" F] by (simp add: frequently_def) lemma shows not_eventually: "\ eventually P F \ (\\<^sub>Fx in F. \ P x)" and not_frequently: "\ frequently P F \ (\\<^sub>Fx in F. \ P x)" by (auto simp: frequently_def) lemma frequently_imp_iff: "(\\<^sub>Fx in F. P x \ Q x) \ (eventually P F \ frequently Q F)" unfolding imp_conv_disj frequently_disj_iff not_eventually[symmetric] .. lemma eventually_frequently_const_simps: "(\\<^sub>Fx in F. P x \ C) \ (\\<^sub>Fx in F. P x) \ C" "(\\<^sub>Fx in F. C \ P x) \ C \ (\\<^sub>Fx in F. P x)" "(\\<^sub>Fx in F. P x \ C) \ (\\<^sub>Fx in F. P x) \ C" "(\\<^sub>Fx in F. C \ P x) \ C \ (\\<^sub>Fx in F. P x)" "(\\<^sub>Fx in F. P x \ C) \ ((\\<^sub>Fx in F. P x) \ C)" "(\\<^sub>Fx in F. C \ P x) \ (C \ (\\<^sub>Fx in F. P x))" by (cases C; simp add: not_frequently)+ lemmas eventually_frequently_simps = eventually_frequently_const_simps not_eventually eventually_conj_iff eventually_ball_finite_distrib eventually_ex not_frequently frequently_disj_iff frequently_bex_finite_distrib frequently_all frequently_imp_iff ML \ fun eventually_elim_tac facts = CONTEXT_SUBGOAL (fn (goal, i) => fn (ctxt, st) => let val mp_facts = facts RL @{thms eventually_rev_mp} val rule = @{thm eventuallyI} |> fold (fn mp_fact => fn th => th RS mp_fact) mp_facts |> funpow (length facts) (fn th => @{thm impI} RS th) val cases_prop = Thm.prop_of (Rule_Cases.internalize_params (rule RS Goal.init (Thm.cterm_of ctxt goal))) val cases = Rule_Cases.make_common ctxt cases_prop [(("elim", []), [])] in CONTEXT_CASES cases (resolve_tac ctxt [rule] i) (ctxt, st) end) \ method_setup eventually_elim = \ Scan.succeed (fn _ => CONTEXT_METHOD (fn facts => eventually_elim_tac facts 1)) \ "elimination of eventually quantifiers" subsubsection \Finer-than relation\ text \\<^term>\F \ F'\ means that filter \<^term>\F\ is finer than filter \<^term>\F'\.\ instantiation filter :: (type) complete_lattice begin definition le_filter_def: "F \ F' \ (\P. eventually P F' \ eventually P F)" definition "(F :: 'a filter) < F' \ F \ F' \ \ F' \ F" definition "top = Abs_filter (\P. \x. P x)" definition "bot = Abs_filter (\P. True)" definition "sup F F' = Abs_filter (\P. eventually P F \ eventually P F')" definition "inf F F' = Abs_filter (\P. \Q R. eventually Q F \ eventually R F' \ (\x. Q x \ R x \ P x))" definition "Sup S = Abs_filter (\P. \F\S. eventually P F)" definition "Inf S = Sup {F::'a filter. \F'\S. F \ F'}" lemma eventually_top [simp]: "eventually P top \ (\x. P x)" unfolding top_filter_def by (rule eventually_Abs_filter, rule is_filter.intro, auto) lemma eventually_bot [simp]: "eventually P bot" unfolding bot_filter_def by (subst eventually_Abs_filter, rule is_filter.intro, auto) lemma eventually_sup: "eventually P (sup F F') \ eventually P F \ eventually P F'" unfolding sup_filter_def by (rule eventually_Abs_filter, rule is_filter.intro) (auto elim!: eventually_rev_mp) lemma eventually_inf: "eventually P (inf F F') \ (\Q R. eventually Q F \ eventually R F' \ (\x. Q x \ R x \ P x))" unfolding inf_filter_def apply (rule eventually_Abs_filter [OF is_filter.intro]) apply (blast intro: eventually_True) apply (force elim!: eventually_conj)+ done lemma eventually_Sup: "eventually P (Sup S) \ (\F\S. eventually P F)" unfolding Sup_filter_def apply (rule eventually_Abs_filter [OF is_filter.intro]) apply (auto intro: eventually_conj elim!: eventually_rev_mp) done instance proof fix F F' F'' :: "'a filter" and S :: "'a filter set" { show "F < F' \ F \ F' \ \ F' \ F" by (rule less_filter_def) } { show "F \ F" unfolding le_filter_def by simp } { assume "F \ F'" and "F' \ F''" thus "F \ F''" unfolding le_filter_def by simp } { assume "F \ F'" and "F' \ F" thus "F = F'" unfolding le_filter_def filter_eq_iff by fast } { show "inf F F' \ F" and "inf F F' \ F'" unfolding le_filter_def eventually_inf by (auto intro: eventually_True) } { assume "F \ F'" and "F \ F''" thus "F \ inf F' F''" unfolding le_filter_def eventually_inf by (auto intro: eventually_mono [OF eventually_conj]) } { show "F \ sup F F'" and "F' \ sup F F'" unfolding le_filter_def eventually_sup by simp_all } { assume "F \ F''" and "F' \ F''" thus "sup F F' \ F''" unfolding le_filter_def eventually_sup by simp } { assume "F'' \ S" thus "Inf S \ F''" unfolding le_filter_def Inf_filter_def eventually_Sup Ball_def by simp } { assume "\F'. F' \ S \ F \ F'" thus "F \ Inf S" unfolding le_filter_def Inf_filter_def eventually_Sup Ball_def by simp } { assume "F \ S" thus "F \ Sup S" unfolding le_filter_def eventually_Sup by simp } { assume "\F. F \ S \ F \ F'" thus "Sup S \ F'" unfolding le_filter_def eventually_Sup by simp } { show "Inf {} = (top::'a filter)" by (auto simp: top_filter_def Inf_filter_def Sup_filter_def) (metis (full_types) top_filter_def always_eventually eventually_top) } { show "Sup {} = (bot::'a filter)" by (auto simp: bot_filter_def Sup_filter_def) } qed end instance filter :: (type) distrib_lattice proof fix F G H :: "'a filter" show "sup F (inf G H) = inf (sup F G) (sup F H)" proof (rule order.antisym) show "inf (sup F G) (sup F H) \ sup F (inf G H)" unfolding le_filter_def eventually_sup proof safe fix P assume 1: "eventually P F" and 2: "eventually P (inf G H)" from 2 obtain Q R where QR: "eventually Q G" "eventually R H" "\x. Q x \ R x \ P x" by (auto simp: eventually_inf) define Q' where "Q' = (\x. Q x \ P x)" define R' where "R' = (\x. R x \ P x)" from 1 have "eventually Q' F" by (elim eventually_mono) (auto simp: Q'_def) moreover from 1 have "eventually R' F" by (elim eventually_mono) (auto simp: R'_def) moreover from QR(1) have "eventually Q' G" by (elim eventually_mono) (auto simp: Q'_def) moreover from QR(2) have "eventually R' H" by (elim eventually_mono)(auto simp: R'_def) moreover from QR have "P x" if "Q' x" "R' x" for x using that by (auto simp: Q'_def R'_def) ultimately show "eventually P (inf (sup F G) (sup F H))" by (auto simp: eventually_inf eventually_sup) qed qed (auto intro: inf.coboundedI1 inf.coboundedI2) qed lemma filter_leD: "F \ F' \ eventually P F' \ eventually P F" unfolding le_filter_def by simp lemma filter_leI: "(\P. eventually P F' \ eventually P F) \ F \ F'" unfolding le_filter_def by simp lemma eventually_False: "eventually (\x. False) F \ F = bot" unfolding filter_eq_iff by (auto elim: eventually_rev_mp) lemma eventually_frequently: "F \ bot \ eventually P F \ frequently P F" using eventually_conj[of P F "\x. \ P x"] by (auto simp add: frequently_def eventually_False) lemma eventually_frequentlyE: assumes "eventually P F" assumes "eventually (\x. \ P x \ Q x) F" "F\bot" shows "frequently Q F" proof - have "eventually Q F" using eventually_conj[OF assms(1,2),simplified] by (auto elim:eventually_mono) then show ?thesis using eventually_frequently[OF \F\bot\] by auto qed lemma eventually_const_iff: "eventually (\x. P) F \ P \ F = bot" by (cases P) (auto simp: eventually_False) lemma eventually_const[simp]: "F \ bot \ eventually (\x. P) F \ P" by (simp add: eventually_const_iff) lemma frequently_const_iff: "frequently (\x. P) F \ P \ F \ bot" by (simp add: frequently_def eventually_const_iff) lemma frequently_const[simp]: "F \ bot \ frequently (\x. P) F \ P" by (simp add: frequently_const_iff) lemma eventually_happens: "eventually P net \ net = bot \ (\x. P x)" by (metis frequentlyE eventually_frequently) lemma eventually_happens': assumes "F \ bot" "eventually P F" shows "\x. P x" using assms eventually_frequently frequentlyE by blast abbreviation (input) trivial_limit :: "'a filter \ bool" where "trivial_limit F \ F = bot" lemma trivial_limit_def: "trivial_limit F \ eventually (\x. False) F" by (rule eventually_False [symmetric]) lemma False_imp_not_eventually: "(\x. \ P x ) \ \ trivial_limit net \ \ eventually (\x. P x) net" by (simp add: eventually_False) lemma eventually_Inf: "eventually P (Inf B) \ (\X\B. finite X \ eventually P (Inf X))" proof - let ?F = "\P. \X\B. finite X \ eventually P (Inf X)" - { fix P have "eventually P (Abs_filter ?F) \ ?F P" - proof (rule eventually_Abs_filter is_filter.intro)+ - show "?F (\x. True)" - by (rule exI[of _ "{}"]) (simp add: le_fun_def) - next - fix P Q - assume "?F P" then guess X .. - moreover - assume "?F Q" then guess Y .. - ultimately show "?F (\x. P x \ Q x)" - by (intro exI[of _ "X \ Y"]) - (auto simp: Inf_union_distrib eventually_inf) - next - fix P Q - assume "?F P" then guess X .. - moreover assume "\x. P x \ Q x" - ultimately show "?F Q" - by (intro exI[of _ X]) (auto elim: eventually_mono) - qed } - note eventually_F = this + have eventually_F: "eventually P (Abs_filter ?F) \ ?F P" for P + proof (rule eventually_Abs_filter is_filter.intro)+ + show "?F (\x. True)" + by (rule exI[of _ "{}"]) (simp add: le_fun_def) + next + fix P Q + assume "?F P" "?F Q" + then obtain X Y where + "X \ B" "finite X" "eventually P (\ X)" + "Y \ B" "finite Y" "eventually Q (\ Y)" by blast + then show "?F (\x. P x \ Q x)" + by (intro exI[of _ "X \ Y"]) (auto simp: Inf_union_distrib eventually_inf) + next + fix P Q + assume "?F P" + then obtain X where "X \ B" "finite X" "eventually P (\ X)" + by blast + moreover assume "\x. P x \ Q x" + ultimately show "?F Q" + by (intro exI[of _ X]) (auto elim: eventually_mono) + qed have "Inf B = Abs_filter ?F" proof (intro antisym Inf_greatest) show "Inf B \ Abs_filter ?F" by (auto simp: le_filter_def eventually_F dest: Inf_superset_mono) next fix F assume "F \ B" then show "Abs_filter ?F \ F" by (auto simp add: le_filter_def eventually_F intro!: exI[of _ "{F}"]) qed then show ?thesis by (simp add: eventually_F) qed lemma eventually_INF: "eventually P (\b\B. F b) \ (\X\B. finite X \ eventually P (\b\X. F b))" unfolding eventually_Inf [of P "F`B"] by (metis finite_imageI image_mono finite_subset_image) lemma Inf_filter_not_bot: fixes B :: "'a filter set" shows "(\X. X \ B \ finite X \ Inf X \ bot) \ Inf B \ bot" unfolding trivial_limit_def eventually_Inf[of _ B] bot_bool_def [symmetric] bot_fun_def [symmetric] bot_unique by simp lemma INF_filter_not_bot: fixes F :: "'i \ 'a filter" shows "(\X. X \ B \ finite X \ (\b\X. F b) \ bot) \ (\b\B. F b) \ bot" unfolding trivial_limit_def eventually_INF [of _ _ B] bot_bool_def [symmetric] bot_fun_def [symmetric] bot_unique by simp lemma eventually_Inf_base: assumes "B \ {}" and base: "\F G. F \ B \ G \ B \ \x\B. x \ inf F G" shows "eventually P (Inf B) \ (\b\B. eventually P b)" proof (subst eventually_Inf, safe) fix X assume "finite X" "X \ B" then have "\b\B. \x\X. b \ x" proof induct case empty then show ?case using \B \ {}\ by auto next case (insert x X) then obtain b where "b \ B" "\x. x \ X \ b \ x" by auto with \insert x X \ B\ base[of b x] show ?case by (auto intro: order_trans) qed then obtain b where "b \ B" "b \ Inf X" by (auto simp: le_Inf_iff) then show "eventually P (Inf X) \ Bex B (eventually P)" by (intro bexI[of _ b]) (auto simp: le_filter_def) qed (auto intro!: exI[of _ "{x}" for x]) lemma eventually_INF_base: "B \ {} \ (\a b. a \ B \ b \ B \ \x\B. F x \ inf (F a) (F b)) \ eventually P (\b\B. F b) \ (\b\B. eventually P (F b))" by (subst eventually_Inf_base) auto lemma eventually_INF1: "i \ I \ eventually P (F i) \ eventually P (\i\I. F i)" using filter_leD[OF INF_lower] . lemma eventually_INF_finite: assumes "finite A" shows "eventually P (\ x\A. F x) \ (\Q. (\x\A. eventually (Q x) (F x)) \ (\y. (\x\A. Q x y) \ P y))" using assms proof (induction arbitrary: P rule: finite_induct) case (insert a A P) from insert.hyps have [simp]: "x \ a" if "x \ A" for x using that by auto have "eventually P (\ x\insert a A. F x) \ (\Q R S. eventually Q (F a) \ (( (\x\A. eventually (S x) (F x)) \ (\y. (\x\A. S x y) \ R y)) \ (\x. Q x \ R x \ P x)))" unfolding ex_simps by (simp add: eventually_inf insert.IH) also have "\ \ (\Q. (\x\insert a A. eventually (Q x) (F x)) \ (\y. (\x\insert a A. Q x y) \ P y))" proof (safe, goal_cases) case (1 Q R S) thus ?case using 1 by (intro exI[of _ "S(a := Q)"]) auto next case (2 Q) show ?case by (rule exI[of _ "Q a"], rule exI[of _ "\y. \x\A. Q x y"], rule exI[of _ "Q(a := (\_. True))"]) (use 2 in auto) qed finally show ?case . qed auto subsubsection \Map function for filters\ definition filtermap :: "('a \ 'b) \ 'a filter \ 'b filter" where "filtermap f F = Abs_filter (\P. eventually (\x. P (f x)) F)" lemma eventually_filtermap: "eventually P (filtermap f F) = eventually (\x. P (f x)) F" unfolding filtermap_def apply (rule eventually_Abs_filter [OF is_filter.intro]) apply (auto elim!: eventually_rev_mp) done lemma filtermap_ident: "filtermap (\x. x) F = F" by (simp add: filter_eq_iff eventually_filtermap) lemma filtermap_filtermap: "filtermap f (filtermap g F) = filtermap (\x. f (g x)) F" by (simp add: filter_eq_iff eventually_filtermap) lemma filtermap_mono: "F \ F' \ filtermap f F \ filtermap f F'" unfolding le_filter_def eventually_filtermap by simp lemma filtermap_bot [simp]: "filtermap f bot = bot" by (simp add: filter_eq_iff eventually_filtermap) lemma filtermap_bot_iff: "filtermap f F = bot \ F = bot" by (simp add: trivial_limit_def eventually_filtermap) lemma filtermap_sup: "filtermap f (sup F1 F2) = sup (filtermap f F1) (filtermap f F2)" by (simp add: filter_eq_iff eventually_filtermap eventually_sup) lemma filtermap_SUP: "filtermap f (\b\B. F b) = (\b\B. filtermap f (F b))" by (simp add: filter_eq_iff eventually_Sup eventually_filtermap) lemma filtermap_inf: "filtermap f (inf F1 F2) \ inf (filtermap f F1) (filtermap f F2)" by (intro inf_greatest filtermap_mono inf_sup_ord) lemma filtermap_INF: "filtermap f (\b\B. F b) \ (\b\B. filtermap f (F b))" by (rule INF_greatest, rule filtermap_mono, erule INF_lower) subsubsection \Contravariant map function for filters\ definition filtercomap :: "('a \ 'b) \ 'b filter \ 'a filter" where "filtercomap f F = Abs_filter (\P. \Q. eventually Q F \ (\x. Q (f x) \ P x))" lemma eventually_filtercomap: "eventually P (filtercomap f F) \ (\Q. eventually Q F \ (\x. Q (f x) \ P x))" unfolding filtercomap_def proof (intro eventually_Abs_filter, unfold_locales, goal_cases) case 1 show ?case by (auto intro!: exI[of _ "\_. True"]) next case (2 P Q) - from 2(1) guess P' by (elim exE conjE) note P' = this - from 2(2) guess Q' by (elim exE conjE) note Q' = this + then obtain P' Q' where P'Q': + "eventually P' F" "\x. P' (f x) \ P x" + "eventually Q' F" "\x. Q' (f x) \ Q x" + by (elim exE conjE) show ?case - by (rule exI[of _ "\x. P' x \ Q' x"]) - (insert P' Q', auto intro!: eventually_conj) + by (rule exI[of _ "\x. P' x \ Q' x"]) (use P'Q' in \auto intro!: eventually_conj\) next case (3 P Q) thus ?case by blast qed lemma filtercomap_ident: "filtercomap (\x. x) F = F" by (auto simp: filter_eq_iff eventually_filtercomap elim!: eventually_mono) lemma filtercomap_filtercomap: "filtercomap f (filtercomap g F) = filtercomap (\x. g (f x)) F" unfolding filter_eq_iff by (auto simp: eventually_filtercomap) lemma filtercomap_mono: "F \ F' \ filtercomap f F \ filtercomap f F'" by (auto simp: eventually_filtercomap le_filter_def) lemma filtercomap_bot [simp]: "filtercomap f bot = bot" by (auto simp: filter_eq_iff eventually_filtercomap) lemma filtercomap_top [simp]: "filtercomap f top = top" by (auto simp: filter_eq_iff eventually_filtercomap) lemma filtercomap_inf: "filtercomap f (inf F1 F2) = inf (filtercomap f F1) (filtercomap f F2)" unfolding filter_eq_iff proof safe fix P assume "eventually P (filtercomap f (F1 \ F2))" then obtain Q R S where *: "eventually Q F1" "eventually R F2" "\x. Q x \ R x \ S x" "\x. S (f x) \ P x" unfolding eventually_filtercomap eventually_inf by blast from * have "eventually (\x. Q (f x)) (filtercomap f F1)" "eventually (\x. R (f x)) (filtercomap f F2)" by (auto simp: eventually_filtercomap) with * show "eventually P (filtercomap f F1 \ filtercomap f F2)" unfolding eventually_inf by blast next fix P assume "eventually P (inf (filtercomap f F1) (filtercomap f F2))" then obtain Q Q' R R' where *: "eventually Q F1" "eventually R F2" "\x. Q (f x) \ Q' x" "\x. R (f x) \ R' x" "\x. Q' x \ R' x \ P x" unfolding eventually_filtercomap eventually_inf by blast from * have "eventually (\x. Q x \ R x) (F1 \ F2)" by (auto simp: eventually_inf) with * show "eventually P (filtercomap f (F1 \ F2))" by (auto simp: eventually_filtercomap) qed lemma filtercomap_sup: "filtercomap f (sup F1 F2) \ sup (filtercomap f F1) (filtercomap f F2)" by (intro sup_least filtercomap_mono inf_sup_ord) lemma filtercomap_INF: "filtercomap f (\b\B. F b) = (\b\B. filtercomap f (F b))" proof - have *: "filtercomap f (\b\B. F b) = (\b\B. filtercomap f (F b))" if "finite B" for B using that by induction (simp_all add: filtercomap_inf) show ?thesis unfolding filter_eq_iff proof fix P have "eventually P (\b\B. filtercomap f (F b)) \ (\X. (X \ B \ finite X) \ eventually P (\b\X. filtercomap f (F b)))" by (subst eventually_INF) blast also have "\ \ (\X. (X \ B \ finite X) \ eventually P (filtercomap f (\b\X. F b)))" by (rule ex_cong) (simp add: *) also have "\ \ eventually P (filtercomap f (\(F ` B)))" unfolding eventually_filtercomap by (subst eventually_INF) blast finally show "eventually P (filtercomap f (\(F ` B))) = eventually P (\b\B. filtercomap f (F b))" .. qed qed lemma filtercomap_SUP: "filtercomap f (\b\B. F b) \ (\b\B. filtercomap f (F b))" by (intro SUP_least filtercomap_mono SUP_upper) lemma filtermap_le_iff_le_filtercomap: "filtermap f F \ G \ F \ filtercomap f G" unfolding le_filter_def eventually_filtermap eventually_filtercomap using eventually_mono by auto lemma filtercomap_neq_bot: assumes "\P. eventually P F \ \x. P (f x)" shows "filtercomap f F \ bot" using assms by (auto simp: trivial_limit_def eventually_filtercomap) lemma filtercomap_neq_bot_surj: assumes "F \ bot" and "surj f" shows "filtercomap f F \ bot" proof (rule filtercomap_neq_bot) fix P assume *: "eventually P F" show "\x. P (f x)" proof (rule ccontr) assume **: "\(\x. P (f x))" from * have "eventually (\_. False) F" proof eventually_elim case (elim x) from \surj f\ obtain y where "x = f y" by auto with elim and ** show False by auto qed with assms show False by (simp add: trivial_limit_def) qed qed lemma eventually_filtercomapI [intro]: assumes "eventually P F" shows "eventually (\x. P (f x)) (filtercomap f F)" using assms by (auto simp: eventually_filtercomap) lemma filtermap_filtercomap: "filtermap f (filtercomap f F) \ F" by (auto simp: le_filter_def eventually_filtermap eventually_filtercomap) lemma filtercomap_filtermap: "filtercomap f (filtermap f F) \ F" unfolding le_filter_def eventually_filtermap eventually_filtercomap by (auto elim!: eventually_mono) subsubsection \Standard filters\ definition principal :: "'a set \ 'a filter" where "principal S = Abs_filter (\P. \x\S. P x)" lemma eventually_principal: "eventually P (principal S) \ (\x\S. P x)" unfolding principal_def by (rule eventually_Abs_filter, rule is_filter.intro) auto lemma eventually_inf_principal: "eventually P (inf F (principal s)) \ eventually (\x. x \ s \ P x) F" unfolding eventually_inf eventually_principal by (auto elim: eventually_mono) lemma principal_UNIV[simp]: "principal UNIV = top" by (auto simp: filter_eq_iff eventually_principal) lemma principal_empty[simp]: "principal {} = bot" by (auto simp: filter_eq_iff eventually_principal) lemma principal_eq_bot_iff: "principal X = bot \ X = {}" by (auto simp add: filter_eq_iff eventually_principal) lemma principal_le_iff[iff]: "principal A \ principal B \ A \ B" by (auto simp: le_filter_def eventually_principal) lemma le_principal: "F \ principal A \ eventually (\x. x \ A) F" unfolding le_filter_def eventually_principal by (force elim: eventually_mono) lemma principal_inject[iff]: "principal A = principal B \ A = B" unfolding eq_iff by simp lemma sup_principal[simp]: "sup (principal A) (principal B) = principal (A \ B)" unfolding filter_eq_iff eventually_sup eventually_principal by auto lemma inf_principal[simp]: "inf (principal A) (principal B) = principal (A \ B)" unfolding filter_eq_iff eventually_inf eventually_principal by (auto intro: exI[of _ "\x. x \ A"] exI[of _ "\x. x \ B"]) lemma SUP_principal[simp]: "(\i\I. principal (A i)) = principal (\i\I. A i)" unfolding filter_eq_iff eventually_Sup by (auto simp: eventually_principal) lemma INF_principal_finite: "finite X \ (\x\X. principal (f x)) = principal (\x\X. f x)" by (induct X rule: finite_induct) auto lemma filtermap_principal[simp]: "filtermap f (principal A) = principal (f ` A)" unfolding filter_eq_iff eventually_filtermap eventually_principal by simp lemma filtercomap_principal[simp]: "filtercomap f (principal A) = principal (f -` A)" unfolding filter_eq_iff eventually_filtercomap eventually_principal by fast subsubsection \Order filters\ definition at_top :: "('a::order) filter" where "at_top = (\k. principal {k ..})" lemma at_top_sub: "at_top = (\k\{c::'a::linorder..}. principal {k ..})" by (auto intro!: INF_eq max.cobounded1 max.cobounded2 simp: at_top_def) lemma eventually_at_top_linorder: "eventually P at_top \ (\N::'a::linorder. \n\N. P n)" unfolding at_top_def by (subst eventually_INF_base) (auto simp: eventually_principal intro: max.cobounded1 max.cobounded2) lemma eventually_filtercomap_at_top_linorder: "eventually P (filtercomap f at_top) \ (\N::'a::linorder. \x. f x \ N \ P x)" by (auto simp: eventually_filtercomap eventually_at_top_linorder) lemma eventually_at_top_linorderI: fixes c::"'a::linorder" assumes "\x. c \ x \ P x" shows "eventually P at_top" using assms by (auto simp: eventually_at_top_linorder) lemma eventually_ge_at_top [simp]: "eventually (\x. (c::_::linorder) \ x) at_top" unfolding eventually_at_top_linorder by auto lemma eventually_at_top_dense: "eventually P at_top \ (\N::'a::{no_top, linorder}. \n>N. P n)" proof - have "eventually P (\k. principal {k <..}) \ (\N::'a. \n>N. P n)" by (subst eventually_INF_base) (auto simp: eventually_principal intro: max.cobounded1 max.cobounded2) also have "(\k. principal {k::'a <..}) = at_top" unfolding at_top_def by (intro INF_eq) (auto intro: less_imp_le simp: Ici_subset_Ioi_iff gt_ex) finally show ?thesis . qed lemma eventually_filtercomap_at_top_dense: "eventually P (filtercomap f at_top) \ (\N::'a::{no_top, linorder}. \x. f x > N \ P x)" by (auto simp: eventually_filtercomap eventually_at_top_dense) lemma eventually_at_top_not_equal [simp]: "eventually (\x::'a::{no_top, linorder}. x \ c) at_top" unfolding eventually_at_top_dense by auto lemma eventually_gt_at_top [simp]: "eventually (\x. (c::_::{no_top, linorder}) < x) at_top" unfolding eventually_at_top_dense by auto lemma eventually_all_ge_at_top: assumes "eventually P (at_top :: ('a :: linorder) filter)" shows "eventually (\x. \y\x. P y) at_top" proof - from assms obtain x where "\y. y \ x \ P y" by (auto simp: eventually_at_top_linorder) hence "\z\y. P z" if "y \ x" for y using that by simp thus ?thesis by (auto simp: eventually_at_top_linorder) qed definition at_bot :: "('a::order) filter" where "at_bot = (\k. principal {.. k})" lemma at_bot_sub: "at_bot = (\k\{.. c::'a::linorder}. principal {.. k})" by (auto intro!: INF_eq min.cobounded1 min.cobounded2 simp: at_bot_def) lemma eventually_at_bot_linorder: fixes P :: "'a::linorder \ bool" shows "eventually P at_bot \ (\N. \n\N. P n)" unfolding at_bot_def by (subst eventually_INF_base) (auto simp: eventually_principal intro: min.cobounded1 min.cobounded2) lemma eventually_filtercomap_at_bot_linorder: "eventually P (filtercomap f at_bot) \ (\N::'a::linorder. \x. f x \ N \ P x)" by (auto simp: eventually_filtercomap eventually_at_bot_linorder) lemma eventually_le_at_bot [simp]: "eventually (\x. x \ (c::_::linorder)) at_bot" unfolding eventually_at_bot_linorder by auto lemma eventually_at_bot_dense: "eventually P at_bot \ (\N::'a::{no_bot, linorder}. \nk. principal {..< k}) \ (\N::'a. \nk. principal {..< k::'a}) = at_bot" unfolding at_bot_def by (intro INF_eq) (auto intro: less_imp_le simp: Iic_subset_Iio_iff lt_ex) finally show ?thesis . qed lemma eventually_filtercomap_at_bot_dense: "eventually P (filtercomap f at_bot) \ (\N::'a::{no_bot, linorder}. \x. f x < N \ P x)" by (auto simp: eventually_filtercomap eventually_at_bot_dense) lemma eventually_at_bot_not_equal [simp]: "eventually (\x::'a::{no_bot, linorder}. x \ c) at_bot" unfolding eventually_at_bot_dense by auto lemma eventually_gt_at_bot [simp]: "eventually (\x. x < (c::_::unbounded_dense_linorder)) at_bot" unfolding eventually_at_bot_dense by auto lemma trivial_limit_at_bot_linorder [simp]: "\ trivial_limit (at_bot ::('a::linorder) filter)" unfolding trivial_limit_def by (metis eventually_at_bot_linorder order_refl) lemma trivial_limit_at_top_linorder [simp]: "\ trivial_limit (at_top ::('a::linorder) filter)" unfolding trivial_limit_def by (metis eventually_at_top_linorder order_refl) subsection \Sequentially\ abbreviation sequentially :: "nat filter" where "sequentially \ at_top" lemma eventually_sequentially: "eventually P sequentially \ (\N. \n\N. P n)" by (rule eventually_at_top_linorder) lemma sequentially_bot [simp, intro]: "sequentially \ bot" unfolding filter_eq_iff eventually_sequentially by auto lemmas trivial_limit_sequentially = sequentially_bot lemma eventually_False_sequentially [simp]: "\ eventually (\n. False) sequentially" by (simp add: eventually_False) lemma le_sequentially: "F \ sequentially \ (\N. eventually (\n. N \ n) F)" by (simp add: at_top_def le_INF_iff le_principal) lemma eventually_sequentiallyI [intro?]: assumes "\x. c \ x \ P x" shows "eventually P sequentially" using assms by (auto simp: eventually_sequentially) lemma eventually_sequentially_Suc [simp]: "eventually (\i. P (Suc i)) sequentially \ eventually P sequentially" unfolding eventually_sequentially by (metis Suc_le_D Suc_le_mono le_Suc_eq) lemma eventually_sequentially_seg [simp]: "eventually (\n. P (n + k)) sequentially \ eventually P sequentially" using eventually_sequentially_Suc[of "\n. P (n + k)" for k] by (induction k) auto lemma filtermap_sequentually_ne_bot: "filtermap f sequentially \ bot" by (simp add: filtermap_bot_iff) subsection \Increasing finite subsets\ definition finite_subsets_at_top where "finite_subsets_at_top A = (\ X\{X. finite X \ X \ A}. principal {Y. finite Y \ X \ Y \ Y \ A})" lemma eventually_finite_subsets_at_top: "eventually P (finite_subsets_at_top A) \ (\X. finite X \ X \ A \ (\Y. finite Y \ X \ Y \ Y \ A \ P Y))" unfolding finite_subsets_at_top_def proof (subst eventually_INF_base, goal_cases) show "{X. finite X \ X \ A} \ {}" by auto next case (2 B C) thus ?case by (intro bexI[of _ "B \ C"]) auto qed (simp_all add: eventually_principal) lemma eventually_finite_subsets_at_top_weakI [intro]: assumes "\X. finite X \ X \ A \ P X" shows "eventually P (finite_subsets_at_top A)" proof - have "eventually (\X. finite X \ X \ A) (finite_subsets_at_top A)" by (auto simp: eventually_finite_subsets_at_top) thus ?thesis by eventually_elim (use assms in auto) qed lemma finite_subsets_at_top_neq_bot [simp]: "finite_subsets_at_top A \ bot" proof - have "\eventually (\x. False) (finite_subsets_at_top A)" by (auto simp: eventually_finite_subsets_at_top) thus ?thesis by auto qed lemma filtermap_image_finite_subsets_at_top: assumes "inj_on f A" shows "filtermap ((`) f) (finite_subsets_at_top A) = finite_subsets_at_top (f ` A)" unfolding filter_eq_iff eventually_filtermap proof (safe, goal_cases) case (1 P) then obtain X where X: "finite X" "X \ A" "\Y. finite Y \ X \ Y \ Y \ A \ P (f ` Y)" unfolding eventually_finite_subsets_at_top by force show ?case unfolding eventually_finite_subsets_at_top eventually_filtermap proof (rule exI[of _ "f ` X"], intro conjI allI impI, goal_cases) case (3 Y) with assms and X(1,2) have "P (f ` (f -` Y \ A))" using X(1,2) by (intro X(3) finite_vimage_IntI) auto also have "f ` (f -` Y \ A) = Y" using assms 3 by blast finally show ?case . qed (insert assms X(1,2), auto intro!: finite_vimage_IntI) next case (2 P) then obtain X where X: "finite X" "X \ f ` A" "\Y. finite Y \ X \ Y \ Y \ f ` A \ P Y" unfolding eventually_finite_subsets_at_top by force show ?case unfolding eventually_finite_subsets_at_top eventually_filtermap proof (rule exI[of _ "f -` X \ A"], intro conjI allI impI, goal_cases) case (3 Y) with X(1,2) and assms show ?case by (intro X(3)) force+ qed (insert assms X(1), auto intro!: finite_vimage_IntI) qed lemma eventually_finite_subsets_at_top_finite: assumes "finite A" shows "eventually P (finite_subsets_at_top A) \ P A" unfolding eventually_finite_subsets_at_top using assms by force lemma finite_subsets_at_top_finite: "finite A \ finite_subsets_at_top A = principal {A}" by (auto simp: filter_eq_iff eventually_finite_subsets_at_top_finite eventually_principal) subsection \The cofinite filter\ definition "cofinite = Abs_filter (\P. finite {x. \ P x})" abbreviation Inf_many :: "('a \ bool) \ bool" (binder "\\<^sub>\" 10) where "Inf_many P \ frequently P cofinite" abbreviation Alm_all :: "('a \ bool) \ bool" (binder "\\<^sub>\" 10) where "Alm_all P \ eventually P cofinite" notation (ASCII) Inf_many (binder "INFM " 10) and Alm_all (binder "MOST " 10) lemma eventually_cofinite: "eventually P cofinite \ finite {x. \ P x}" unfolding cofinite_def proof (rule eventually_Abs_filter, rule is_filter.intro) fix P Q :: "'a \ bool" assume "finite {x. \ P x}" "finite {x. \ Q x}" from finite_UnI[OF this] show "finite {x. \ (P x \ Q x)}" by (rule rev_finite_subset) auto next fix P Q :: "'a \ bool" assume P: "finite {x. \ P x}" and *: "\x. P x \ Q x" from * show "finite {x. \ Q x}" by (intro finite_subset[OF _ P]) auto qed simp lemma frequently_cofinite: "frequently P cofinite \ \ finite {x. P x}" by (simp add: frequently_def eventually_cofinite) lemma cofinite_bot[simp]: "cofinite = (bot::'a filter) \ finite (UNIV :: 'a set)" unfolding trivial_limit_def eventually_cofinite by simp lemma cofinite_eq_sequentially: "cofinite = sequentially" unfolding filter_eq_iff eventually_sequentially eventually_cofinite proof safe fix P :: "nat \ bool" assume [simp]: "finite {x. \ P x}" show "\N. \n\N. P n" proof cases assume "{x. \ P x} \ {}" then show ?thesis by (intro exI[of _ "Suc (Max {x. \ P x})"]) (auto simp: Suc_le_eq) qed auto next fix P :: "nat \ bool" and N :: nat assume "\n\N. P n" then have "{x. \ P x} \ {..< N}" by (auto simp: not_le) then show "finite {x. \ P x}" by (blast intro: finite_subset) qed subsubsection \Product of filters\ definition prod_filter :: "'a filter \ 'b filter \ ('a \ 'b) filter" (infixr "\\<^sub>F" 80) where "prod_filter F G = (\(P, Q)\{(P, Q). eventually P F \ eventually Q G}. principal {(x, y). P x \ Q y})" lemma eventually_prod_filter: "eventually P (F \\<^sub>F G) \ (\Pf Pg. eventually Pf F \ eventually Pg G \ (\x y. Pf x \ Pg y \ P (x, y)))" unfolding prod_filter_def proof (subst eventually_INF_base, goal_cases) case 2 moreover have "eventually Pf F \ eventually Qf F \ eventually Pg G \ eventually Qg G \ \P Q. eventually P F \ eventually Q G \ Collect P \ Collect Q \ Collect Pf \ Collect Pg \ Collect Qf \ Collect Qg" for Pf Pg Qf Qg by (intro conjI exI[of _ "inf Pf Qf"] exI[of _ "inf Pg Qg"]) (auto simp: inf_fun_def eventually_conj) ultimately show ?case by auto qed (auto simp: eventually_principal intro: eventually_True) lemma eventually_prod1: assumes "B \ bot" shows "(\\<^sub>F (x, y) in A \\<^sub>F B. P x) \ (\\<^sub>F x in A. P x)" unfolding eventually_prod_filter proof safe fix R Q assume *: "\\<^sub>F x in A. R x" "\\<^sub>F x in B. Q x" "\x y. R x \ Q y \ P x" with \B \ bot\ obtain y where "Q y" by (auto dest: eventually_happens) with * show "eventually P A" by (force elim: eventually_mono) next assume "eventually P A" then show "\Pf Pg. eventually Pf A \ eventually Pg B \ (\x y. Pf x \ Pg y \ P x)" by (intro exI[of _ P] exI[of _ "\x. True"]) auto qed lemma eventually_prod2: assumes "A \ bot" shows "(\\<^sub>F (x, y) in A \\<^sub>F B. P y) \ (\\<^sub>F y in B. P y)" unfolding eventually_prod_filter proof safe fix R Q assume *: "\\<^sub>F x in A. R x" "\\<^sub>F x in B. Q x" "\x y. R x \ Q y \ P y" with \A \ bot\ obtain x where "R x" by (auto dest: eventually_happens) with * show "eventually P B" by (force elim: eventually_mono) next assume "eventually P B" then show "\Pf Pg. eventually Pf A \ eventually Pg B \ (\x y. Pf x \ Pg y \ P y)" by (intro exI[of _ P] exI[of _ "\x. True"]) auto qed lemma INF_filter_bot_base: fixes F :: "'a \ 'b filter" assumes *: "\i j. i \ I \ j \ I \ \k\I. F k \ F i \ F j" shows "(\i\I. F i) = bot \ (\i\I. F i = bot)" proof (cases "\i\I. F i = bot") case True then have "(\i\I. F i) \ bot" by (auto intro: INF_lower2) with True show ?thesis by (auto simp: bot_unique) next case False moreover have "(\i\I. F i) \ bot" proof (cases "I = {}") case True then show ?thesis by (auto simp add: filter_eq_iff) next case False': False show ?thesis proof (rule INF_filter_not_bot) fix J assume "finite J" "J \ I" then have "\k\I. F k \ (\i\J. F i)" proof (induct J) case empty then show ?case using \I \ {}\ by auto next case (insert i J) then obtain k where "k \ I" "F k \ (\i\J. F i)" by auto with insert *[of i k] show ?case by auto qed with False show "(\i\J. F i) \ \" by (auto simp: bot_unique) qed qed ultimately show ?thesis by auto qed lemma Collect_empty_eq_bot: "Collect P = {} \ P = \" by auto lemma prod_filter_eq_bot: "A \\<^sub>F B = bot \ A = bot \ B = bot" unfolding trivial_limit_def proof assume "\\<^sub>F x in A \\<^sub>F B. False" then obtain Pf Pg where Pf: "eventually (\x. Pf x) A" and Pg: "eventually (\y. Pg y) B" and *: "\x y. Pf x \ Pg y \ False" unfolding eventually_prod_filter by fast from * have "(\x. \ Pf x) \ (\y. \ Pg y)" by fast with Pf Pg show "(\\<^sub>F x in A. False) \ (\\<^sub>F x in B. False)" by auto next assume "(\\<^sub>F x in A. False) \ (\\<^sub>F x in B. False)" then show "\\<^sub>F x in A \\<^sub>F B. False" unfolding eventually_prod_filter by (force intro: eventually_True) qed lemma prod_filter_mono: "F \ F' \ G \ G' \ F \\<^sub>F G \ F' \\<^sub>F G'" by (auto simp: le_filter_def eventually_prod_filter) lemma prod_filter_mono_iff: assumes nAB: "A \ bot" "B \ bot" shows "A \\<^sub>F B \ C \\<^sub>F D \ A \ C \ B \ D" proof safe assume *: "A \\<^sub>F B \ C \\<^sub>F D" with assms have "A \\<^sub>F B \ bot" by (auto simp: bot_unique prod_filter_eq_bot) with * have "C \\<^sub>F D \ bot" by (auto simp: bot_unique) then have nCD: "C \ bot" "D \ bot" by (auto simp: prod_filter_eq_bot) show "A \ C" proof (rule filter_leI) fix P assume "eventually P C" with *[THEN filter_leD, of "\(x, y). P x"] show "eventually P A" using nAB nCD by (simp add: eventually_prod1 eventually_prod2) qed show "B \ D" proof (rule filter_leI) fix P assume "eventually P D" with *[THEN filter_leD, of "\(x, y). P y"] show "eventually P B" using nAB nCD by (simp add: eventually_prod1 eventually_prod2) qed qed (intro prod_filter_mono) lemma eventually_prod_same: "eventually P (F \\<^sub>F F) \ (\Q. eventually Q F \ (\x y. Q x \ Q y \ P (x, y)))" unfolding eventually_prod_filter by (blast intro!: eventually_conj) lemma eventually_prod_sequentially: "eventually P (sequentially \\<^sub>F sequentially) \ (\N. \m \ N. \n \ N. P (n, m))" unfolding eventually_prod_same eventually_sequentially by auto lemma principal_prod_principal: "principal A \\<^sub>F principal B = principal (A \ B)" unfolding filter_eq_iff eventually_prod_filter eventually_principal by (fast intro: exI[of _ "\x. x \ A"] exI[of _ "\x. x \ B"]) lemma le_prod_filterI: "filtermap fst F \ A \ filtermap snd F \ B \ F \ A \\<^sub>F B" unfolding le_filter_def eventually_filtermap eventually_prod_filter by (force elim: eventually_elim2) lemma filtermap_fst_prod_filter: "filtermap fst (A \\<^sub>F B) \ A" unfolding le_filter_def eventually_filtermap eventually_prod_filter by (force intro: eventually_True) lemma filtermap_snd_prod_filter: "filtermap snd (A \\<^sub>F B) \ B" unfolding le_filter_def eventually_filtermap eventually_prod_filter by (force intro: eventually_True) lemma prod_filter_INF: assumes "I \ {}" and "J \ {}" shows "(\i\I. A i) \\<^sub>F (\j\J. B j) = (\i\I. \j\J. A i \\<^sub>F B j)" proof (rule antisym) from \I \ {}\ obtain i where "i \ I" by auto from \J \ {}\ obtain j where "j \ J" by auto show "(\i\I. \j\J. A i \\<^sub>F B j) \ (\i\I. A i) \\<^sub>F (\j\J. B j)" by (fast intro: le_prod_filterI INF_greatest INF_lower2 order_trans[OF filtermap_INF] \i \ I\ \j \ J\ filtermap_fst_prod_filter filtermap_snd_prod_filter) show "(\i\I. A i) \\<^sub>F (\j\J. B j) \ (\i\I. \j\J. A i \\<^sub>F B j)" by (intro INF_greatest prod_filter_mono INF_lower) qed lemma filtermap_Pair: "filtermap (\x. (f x, g x)) F \ filtermap f F \\<^sub>F filtermap g F" by (rule le_prod_filterI, simp_all add: filtermap_filtermap) lemma eventually_prodI: "eventually P F \ eventually Q G \ eventually (\x. P (fst x) \ Q (snd x)) (F \\<^sub>F G)" unfolding eventually_prod_filter by auto lemma prod_filter_INF1: "I \ {} \ (\i\I. A i) \\<^sub>F B = (\i\I. A i \\<^sub>F B)" using prod_filter_INF[of I "{B}" A "\x. x"] by simp lemma prod_filter_INF2: "J \ {} \ A \\<^sub>F (\i\J. B i) = (\i\J. A \\<^sub>F B i)" using prod_filter_INF[of "{A}" J "\x. x" B] by simp lemma prod_filtermap1: "prod_filter (filtermap f F) G = filtermap (apfst f) (prod_filter F G)" unfolding filter_eq_iff eventually_filtermap eventually_prod_filter apply safe subgoal by auto subgoal for P Q R by(rule exI[where x="\y. \x. y = f x \ Q x"])(auto intro: eventually_mono) done lemma prod_filtermap2: "prod_filter F (filtermap g G) = filtermap (apsnd g) (prod_filter F G)" unfolding filter_eq_iff eventually_filtermap eventually_prod_filter apply safe subgoal by auto subgoal for P Q R by(auto intro: exI[where x="\y. \x. y = g x \ R x"] eventually_mono) done lemma prod_filter_assoc: "prod_filter (prod_filter F G) H = filtermap (\(x, y, z). ((x, y), z)) (prod_filter F (prod_filter G H))" apply(clarsimp simp add: filter_eq_iff eventually_filtermap eventually_prod_filter; safe) subgoal for P Q R S T by(auto 4 4 intro: exI[where x="\(a, b). T a \ S b"]) subgoal for P Q R S T by(auto 4 3 intro: exI[where x="\(a, b). Q a \ S b"]) done lemma prod_filter_principal_singleton: "prod_filter (principal {x}) F = filtermap (Pair x) F" by(fastforce simp add: filter_eq_iff eventually_prod_filter eventually_principal eventually_filtermap elim: eventually_mono intro: exI[where x="\a. a = x"]) lemma prod_filter_principal_singleton2: "prod_filter F (principal {x}) = filtermap (\a. (a, x)) F" by(fastforce simp add: filter_eq_iff eventually_prod_filter eventually_principal eventually_filtermap elim: eventually_mono intro: exI[where x="\a. a = x"]) lemma prod_filter_commute: "prod_filter F G = filtermap prod.swap (prod_filter G F)" by(auto simp add: filter_eq_iff eventually_prod_filter eventually_filtermap) subsection \Limits\ definition filterlim :: "('a \ 'b) \ 'b filter \ 'a filter \ bool" where "filterlim f F2 F1 \ filtermap f F1 \ F2" syntax "_LIM" :: "pttrns \ 'a \ 'b \ 'a \ bool" ("(3LIM (_)/ (_)./ (_) :> (_))" [1000, 10, 0, 10] 10) translations "LIM x F1. f :> F2" == "CONST filterlim (\x. f) F2 F1" lemma filterlim_top [simp]: "filterlim f top F" by (simp add: filterlim_def) lemma filterlim_iff: "(LIM x F1. f x :> F2) \ (\P. eventually P F2 \ eventually (\x. P (f x)) F1)" unfolding filterlim_def le_filter_def eventually_filtermap .. lemma filterlim_compose: "filterlim g F3 F2 \ filterlim f F2 F1 \ filterlim (\x. g (f x)) F3 F1" unfolding filterlim_def filtermap_filtermap[symmetric] by (metis filtermap_mono order_trans) lemma filterlim_mono: "filterlim f F2 F1 \ F2 \ F2' \ F1' \ F1 \ filterlim f F2' F1'" unfolding filterlim_def by (metis filtermap_mono order_trans) lemma filterlim_ident: "LIM x F. x :> F" by (simp add: filterlim_def filtermap_ident) lemma filterlim_cong: "F1 = F1' \ F2 = F2' \ eventually (\x. f x = g x) F2 \ filterlim f F1 F2 = filterlim g F1' F2'" by (auto simp: filterlim_def le_filter_def eventually_filtermap elim: eventually_elim2) lemma filterlim_mono_eventually: assumes "filterlim f F G" and ord: "F \ F'" "G' \ G" assumes eq: "eventually (\x. f x = f' x) G'" shows "filterlim f' F' G'" proof - have "filterlim f F' G'" by (simp add: filterlim_mono[OF _ ord] assms) then show ?thesis by (rule filterlim_cong[OF refl refl eq, THEN iffD1]) qed lemma filtermap_mono_strong: "inj f \ filtermap f F \ filtermap f G \ F \ G" apply (safe intro!: filtermap_mono) apply (auto simp: le_filter_def eventually_filtermap) apply (erule_tac x="\x. P (inv f x)" in allE) apply auto done lemma eventually_compose_filterlim: assumes "eventually P F" "filterlim f F G" shows "eventually (\x. P (f x)) G" using assms by (simp add: filterlim_iff) lemma filtermap_eq_strong: "inj f \ filtermap f F = filtermap f G \ F = G" by (simp add: filtermap_mono_strong eq_iff) lemma filtermap_fun_inverse: assumes g: "filterlim g F G" assumes f: "filterlim f G F" assumes ev: "eventually (\x. f (g x) = x) G" shows "filtermap f F = G" proof (rule antisym) show "filtermap f F \ G" using f unfolding filterlim_def . have "G = filtermap f (filtermap g G)" using ev by (auto elim: eventually_elim2 simp: filter_eq_iff eventually_filtermap) also have "\ \ filtermap f F" using g by (intro filtermap_mono) (simp add: filterlim_def) finally show "G \ filtermap f F" . qed lemma filterlim_principal: "(LIM x F. f x :> principal S) \ (eventually (\x. f x \ S) F)" unfolding filterlim_def eventually_filtermap le_principal .. lemma filterlim_filtercomap [intro]: "filterlim f F (filtercomap f F)" unfolding filterlim_def by (rule filtermap_filtercomap) lemma filterlim_inf: "(LIM x F1. f x :> inf F2 F3) \ ((LIM x F1. f x :> F2) \ (LIM x F1. f x :> F3))" unfolding filterlim_def by simp lemma filterlim_INF: "(LIM x F. f x :> (\b\B. G b)) \ (\b\B. LIM x F. f x :> G b)" unfolding filterlim_def le_INF_iff .. lemma filterlim_INF_INF: "(\m. m \ J \ \i\I. filtermap f (F i) \ G m) \ LIM x (\i\I. F i). f x :> (\j\J. G j)" unfolding filterlim_def by (rule order_trans[OF filtermap_INF INF_mono]) lemma filterlim_INF': "x \ A \ filterlim f F (G x) \ filterlim f F (\ x\A. G x)" unfolding filterlim_def by (rule order.trans[OF filtermap_mono[OF INF_lower]]) lemma filterlim_filtercomap_iff: "filterlim f (filtercomap g G) F \ filterlim (g \ f) G F" by (simp add: filterlim_def filtermap_le_iff_le_filtercomap filtercomap_filtercomap o_def) lemma filterlim_iff_le_filtercomap: "filterlim f F G \ G \ filtercomap f F" by (simp add: filterlim_def filtermap_le_iff_le_filtercomap) lemma filterlim_base: "(\m x. m \ J \ i m \ I) \ (\m x. m \ J \ x \ F (i m) \ f x \ G m) \ LIM x (\i\I. principal (F i)). f x :> (\j\J. principal (G j))" by (force intro!: filterlim_INF_INF simp: image_subset_iff) lemma filterlim_base_iff: assumes "I \ {}" and chain: "\i j. i \ I \ j \ I \ F i \ F j \ F j \ F i" shows "(LIM x (\i\I. principal (F i)). f x :> \j\J. principal (G j)) \ (\j\J. \i\I. \x\F i. f x \ G j)" unfolding filterlim_INF filterlim_principal proof (subst eventually_INF_base) fix i j assume "i \ I" "j \ I" with chain[OF this] show "\x\I. principal (F x) \ inf (principal (F i)) (principal (F j))" by auto qed (auto simp: eventually_principal \I \ {}\) lemma filterlim_filtermap: "filterlim f F1 (filtermap g F2) = filterlim (\x. f (g x)) F1 F2" unfolding filterlim_def filtermap_filtermap .. lemma filterlim_sup: "filterlim f F F1 \ filterlim f F F2 \ filterlim f F (sup F1 F2)" unfolding filterlim_def filtermap_sup by auto lemma filterlim_sequentially_Suc: "(LIM x sequentially. f (Suc x) :> F) \ (LIM x sequentially. f x :> F)" unfolding filterlim_iff by (subst eventually_sequentially_Suc) simp lemma filterlim_Suc: "filterlim Suc sequentially sequentially" by (simp add: filterlim_iff eventually_sequentially) lemma filterlim_If: "LIM x inf F (principal {x. P x}). f x :> G \ LIM x inf F (principal {x. \ P x}). g x :> G \ LIM x F. if P x then f x else g x :> G" unfolding filterlim_iff eventually_inf_principal by (auto simp: eventually_conj_iff) lemma filterlim_Pair: "LIM x F. f x :> G \ LIM x F. g x :> H \ LIM x F. (f x, g x) :> G \\<^sub>F H" unfolding filterlim_def by (rule order_trans[OF filtermap_Pair prod_filter_mono]) subsection \Limits to \<^const>\at_top\ and \<^const>\at_bot\\ lemma filterlim_at_top: fixes f :: "'a \ ('b::linorder)" shows "(LIM x F. f x :> at_top) \ (\Z. eventually (\x. Z \ f x) F)" by (auto simp: filterlim_iff eventually_at_top_linorder elim!: eventually_mono) lemma filterlim_at_top_mono: "LIM x F. f x :> at_top \ eventually (\x. f x \ (g x::'a::linorder)) F \ LIM x F. g x :> at_top" by (auto simp: filterlim_at_top elim: eventually_elim2 intro: order_trans) lemma filterlim_at_top_dense: fixes f :: "'a \ ('b::unbounded_dense_linorder)" shows "(LIM x F. f x :> at_top) \ (\Z. eventually (\x. Z < f x) F)" by (metis eventually_mono[of _ F] eventually_gt_at_top order_less_imp_le filterlim_at_top[of f F] filterlim_iff[of f at_top F]) lemma filterlim_at_top_ge: fixes f :: "'a \ ('b::linorder)" and c :: "'b" shows "(LIM x F. f x :> at_top) \ (\Z\c. eventually (\x. Z \ f x) F)" unfolding at_top_sub[of c] filterlim_INF by (auto simp add: filterlim_principal) lemma filterlim_at_top_at_top: fixes f :: "'a::linorder \ 'b::linorder" assumes mono: "\x y. Q x \ Q y \ x \ y \ f x \ f y" assumes bij: "\x. P x \ f (g x) = x" "\x. P x \ Q (g x)" assumes Q: "eventually Q at_top" assumes P: "eventually P at_top" shows "filterlim f at_top at_top" proof - from P obtain x where x: "\y. x \ y \ P y" unfolding eventually_at_top_linorder by auto show ?thesis proof (intro filterlim_at_top_ge[THEN iffD2] allI impI) fix z assume "x \ z" with x have "P z" by auto have "eventually (\x. g z \ x) at_top" by (rule eventually_ge_at_top) with Q show "eventually (\x. z \ f x) at_top" by eventually_elim (metis mono bij \P z\) qed qed lemma filterlim_at_top_gt: fixes f :: "'a \ ('b::unbounded_dense_linorder)" and c :: "'b" shows "(LIM x F. f x :> at_top) \ (\Z>c. eventually (\x. Z \ f x) F)" by (metis filterlim_at_top order_less_le_trans gt_ex filterlim_at_top_ge) lemma filterlim_at_bot: fixes f :: "'a \ ('b::linorder)" shows "(LIM x F. f x :> at_bot) \ (\Z. eventually (\x. f x \ Z) F)" by (auto simp: filterlim_iff eventually_at_bot_linorder elim!: eventually_mono) lemma filterlim_at_bot_dense: fixes f :: "'a \ ('b::{dense_linorder, no_bot})" shows "(LIM x F. f x :> at_bot) \ (\Z. eventually (\x. f x < Z) F)" proof (auto simp add: filterlim_at_bot[of f F]) fix Z :: 'b from lt_ex [of Z] obtain Z' where 1: "Z' < Z" .. assume "\Z. eventually (\x. f x \ Z) F" hence "eventually (\x. f x \ Z') F" by auto thus "eventually (\x. f x < Z) F" by (rule eventually_mono) (use 1 in auto) next fix Z :: 'b show "\Z. eventually (\x. f x < Z) F \ eventually (\x. f x \ Z) F" by (drule spec [of _ Z], erule eventually_mono, auto simp add: less_imp_le) qed lemma filterlim_at_bot_le: fixes f :: "'a \ ('b::linorder)" and c :: "'b" shows "(LIM x F. f x :> at_bot) \ (\Z\c. eventually (\x. Z \ f x) F)" unfolding filterlim_at_bot proof safe fix Z assume *: "\Z\c. eventually (\x. Z \ f x) F" with *[THEN spec, of "min Z c"] show "eventually (\x. Z \ f x) F" by (auto elim!: eventually_mono) qed simp lemma filterlim_at_bot_lt: fixes f :: "'a \ ('b::unbounded_dense_linorder)" and c :: "'b" shows "(LIM x F. f x :> at_bot) \ (\Zx. Z \ f x) F)" by (metis filterlim_at_bot filterlim_at_bot_le lt_ex order_le_less_trans) lemma filterlim_finite_subsets_at_top: "filterlim f (finite_subsets_at_top A) F \ (\X. finite X \ X \ A \ eventually (\y. finite (f y) \ X \ f y \ f y \ A) F)" (is "?lhs = ?rhs") proof assume ?lhs thus ?rhs proof (safe, goal_cases) case (1 X) hence *: "(\\<^sub>F x in F. P (f x))" if "eventually P (finite_subsets_at_top A)" for P using that by (auto simp: filterlim_def le_filter_def eventually_filtermap) have "\\<^sub>F Y in finite_subsets_at_top A. finite Y \ X \ Y \ Y \ A" using 1 unfolding eventually_finite_subsets_at_top by force thus ?case by (intro *) auto qed next assume rhs: ?rhs show ?lhs unfolding filterlim_def le_filter_def eventually_finite_subsets_at_top proof (safe, goal_cases) case (1 P X) with rhs have "\\<^sub>F y in F. finite (f y) \ X \ f y \ f y \ A" by auto thus "eventually P (filtermap f F)" unfolding eventually_filtermap by eventually_elim (insert 1, auto) qed qed lemma filterlim_atMost_at_top: "filterlim (\n. {..n}) (finite_subsets_at_top (UNIV :: nat set)) at_top" unfolding filterlim_finite_subsets_at_top proof (safe, goal_cases) case (1 X) then obtain n where n: "X \ {..n}" by (auto simp: finite_nat_set_iff_bounded_le) show ?case using eventually_ge_at_top[of n] by eventually_elim (insert n, auto) qed lemma filterlim_lessThan_at_top: "filterlim (\n. {.. {..Setup \<^typ>\'a filter\ for lifting and transfer\ lemma filtermap_id [simp, id_simps]: "filtermap id = id" by(simp add: fun_eq_iff id_def filtermap_ident) lemma filtermap_id' [simp]: "filtermap (\x. x) = (\F. F)" using filtermap_id unfolding id_def . context includes lifting_syntax begin definition map_filter_on :: "'a set \ ('a \ 'b) \ 'a filter \ 'b filter" where "map_filter_on X f F = Abs_filter (\P. eventually (\x. P (f x) \ x \ X) F)" lemma is_filter_map_filter_on: "is_filter (\P. \\<^sub>F x in F. P (f x) \ x \ X) \ eventually (\x. x \ X) F" proof(rule iffI; unfold_locales) show "\\<^sub>F x in F. True \ x \ X" if "eventually (\x. x \ X) F" using that by simp show "\\<^sub>F x in F. (P (f x) \ Q (f x)) \ x \ X" if "\\<^sub>F x in F. P (f x) \ x \ X" "\\<^sub>F x in F. Q (f x) \ x \ X" for P Q using eventually_conj[OF that] by(auto simp add: conj_ac cong: conj_cong) show "\\<^sub>F x in F. Q (f x) \ x \ X" if "\x. P x \ Q x" "\\<^sub>F x in F. P (f x) \ x \ X" for P Q using that(2) by(rule eventually_mono)(use that(1) in auto) show "eventually (\x. x \ X) F" if "is_filter (\P. \\<^sub>F x in F. P (f x) \ x \ X)" using is_filter.True[OF that] by simp qed lemma eventually_map_filter_on: "eventually P (map_filter_on X f F) = (\\<^sub>F x in F. P (f x) \ x \ X)" if "eventually (\x. x \ X) F" by(simp add: is_filter_map_filter_on map_filter_on_def eventually_Abs_filter that) lemma map_filter_on_UNIV: "map_filter_on UNIV = filtermap" by(simp add: map_filter_on_def filtermap_def fun_eq_iff) lemma map_filter_on_comp: "map_filter_on X f (map_filter_on Y g F) = map_filter_on Y (f \ g) F" if "g ` Y \ X" and "eventually (\x. x \ Y) F" unfolding map_filter_on_def using that(1) by(auto simp add: eventually_Abs_filter that(2) is_filter_map_filter_on intro!: arg_cong[where f=Abs_filter] arg_cong2[where f=eventually]) inductive rel_filter :: "('a \ 'b \ bool) \ 'a filter \ 'b filter \ bool" for R F G where "rel_filter R F G" if "eventually (case_prod R) Z" "map_filter_on {(x, y). R x y} fst Z = F" "map_filter_on {(x, y). R x y} snd Z = G" lemma rel_filter_eq [relator_eq]: "rel_filter (=) = (=)" proof(intro ext iffI)+ show "F = G" if "rel_filter (=) F G" for F G using that by cases(clarsimp simp add: filter_eq_iff eventually_map_filter_on split_def cong: rev_conj_cong) show "rel_filter (=) F G" if "F = G" for F G unfolding \F = G\ proof let ?Z = "map_filter_on UNIV (\x. (x, x)) G" have [simp]: "range (\x. (x, x)) \ {(x, y). x = y}" by auto show "map_filter_on {(x, y). x = y} fst ?Z = G" and "map_filter_on {(x, y). x = y} snd ?Z = G" by(simp_all add: map_filter_on_comp)(simp_all add: map_filter_on_UNIV o_def) show "\\<^sub>F (x, y) in ?Z. x = y" by(simp add: eventually_map_filter_on) qed qed lemma rel_filter_mono [relator_mono]: "rel_filter A \ rel_filter B" if le: "A \ B" proof(clarify elim!: rel_filter.cases) show "rel_filter B (map_filter_on {(x, y). A x y} fst Z) (map_filter_on {(x, y). A x y} snd Z)" (is "rel_filter _ ?X ?Y") if "\\<^sub>F (x, y) in Z. A x y" for Z proof let ?Z = "map_filter_on {(x, y). A x y} id Z" show "\\<^sub>F (x, y) in ?Z. B x y" using le that by(simp add: eventually_map_filter_on le_fun_def split_def conj_commute cong: conj_cong) have [simp]: "{(x, y). A x y} \ {(x, y). B x y}" using le by auto show "map_filter_on {(x, y). B x y} fst ?Z = ?X" "map_filter_on {(x, y). B x y} snd ?Z = ?Y" using le that by(simp_all add: le_fun_def map_filter_on_comp) qed qed lemma rel_filter_conversep: "rel_filter A\\ = (rel_filter A)\\" proof(safe intro!: ext elim!: rel_filter.cases) show *: "rel_filter A (map_filter_on {(x, y). A\\ x y} snd Z) (map_filter_on {(x, y). A\\ x y} fst Z)" (is "rel_filter _ ?X ?Y") if "\\<^sub>F (x, y) in Z. A\\ x y" for A Z proof let ?Z = "map_filter_on {(x, y). A y x} prod.swap Z" show "\\<^sub>F (x, y) in ?Z. A x y" using that by(simp add: eventually_map_filter_on) have [simp]: "prod.swap ` {(x, y). A y x} \ {(x, y). A x y}" by auto show "map_filter_on {(x, y). A x y} fst ?Z = ?X" "map_filter_on {(x, y). A x y} snd ?Z = ?Y" using that by(simp_all add: map_filter_on_comp o_def) qed show "rel_filter A\\ (map_filter_on {(x, y). A x y} snd Z) (map_filter_on {(x, y). A x y} fst Z)" if "\\<^sub>F (x, y) in Z. A x y" for Z using *[of "A\\" Z] that by simp qed lemma rel_filter_distr [relator_distr]: "rel_filter A OO rel_filter B = rel_filter (A OO B)" proof(safe intro!: ext elim!: rel_filter.cases) let ?AB = "{(x, y). (A OO B) x y}" show "(rel_filter A OO rel_filter B) (map_filter_on {(x, y). (A OO B) x y} fst Z) (map_filter_on {(x, y). (A OO B) x y} snd Z)" (is "(_ OO _) ?F ?H") if "\\<^sub>F (x, y) in Z. (A OO B) x y" for Z proof let ?G = "map_filter_on ?AB (\(x, y). SOME z. A x z \ B z y) Z" show "rel_filter A ?F ?G" proof let ?Z = "map_filter_on ?AB (\(x, y). (x, SOME z. A x z \ B z y)) Z" show "\\<^sub>F (x, y) in ?Z. A x y" using that by(auto simp add: eventually_map_filter_on split_def elim!: eventually_mono intro: someI2) have [simp]: "(\p. (fst p, SOME z. A (fst p) z \ B z (snd p))) ` {p. (A OO B) (fst p) (snd p)} \ {p. A (fst p) (snd p)}" by(auto intro: someI2) show "map_filter_on {(x, y). A x y} fst ?Z = ?F" "map_filter_on {(x, y). A x y} snd ?Z = ?G" using that by(simp_all add: map_filter_on_comp split_def o_def) qed show "rel_filter B ?G ?H" proof let ?Z = "map_filter_on ?AB (\(x, y). (SOME z. A x z \ B z y, y)) Z" show "\\<^sub>F (x, y) in ?Z. B x y" using that by(auto simp add: eventually_map_filter_on split_def elim!: eventually_mono intro: someI2) have [simp]: "(\p. (SOME z. A (fst p) z \ B z (snd p), snd p)) ` {p. (A OO B) (fst p) (snd p)} \ {p. B (fst p) (snd p)}" by(auto intro: someI2) show "map_filter_on {(x, y). B x y} fst ?Z = ?G" "map_filter_on {(x, y). B x y} snd ?Z = ?H" using that by(simp_all add: map_filter_on_comp split_def o_def) qed qed fix F G assume F: "\\<^sub>F (x, y) in F. A x y" and G: "\\<^sub>F (x, y) in G. B x y" and eq: "map_filter_on {(x, y). B x y} fst G = map_filter_on {(x, y). A x y} snd F" (is "?Y2 = ?Y1") let ?X = "map_filter_on {(x, y). A x y} fst F" and ?Z = "(map_filter_on {(x, y). B x y} snd G)" have step: "\P'\P. \Q' \ Q. eventually P' F \ eventually Q' G \ {y. \x. P' (x, y)} = {y. \z. Q' (y, z)}" if P: "eventually P F" and Q: "eventually Q G" for P Q proof - let ?P = "\(x, y). P (x, y) \ A x y" and ?Q = "\(y, z). Q (y, z) \ B y z" define P' where "P' \ \(x, y). ?P (x, y) \ (\z. ?Q (y, z))" define Q' where "Q' \ \(y, z). ?Q (y, z) \ (\x. ?P (x, y))" have "P' \ P" "Q' \ Q" "{y. \x. P' (x, y)} = {y. \z. Q' (y, z)}" by(auto simp add: P'_def Q'_def) moreover from P Q F G have P': "eventually ?P F" and Q': "eventually ?Q G" by(simp_all add: eventually_conj_iff split_def) from P' F have "\\<^sub>F y in ?Y1. \x. P (x, y) \ A x y" by(auto simp add: eventually_map_filter_on elim!: eventually_mono) from this[folded eq] obtain Q'' where Q'': "eventually Q'' G" and Q''P: "{y. \z. Q'' (y, z)} \ {y. \x. ?P (x, y)}" using G by(fastforce simp add: eventually_map_filter_on) have "eventually (inf Q'' ?Q) G" using Q'' Q' by(auto intro: eventually_conj simp add: inf_fun_def) then have "eventually Q' G" using Q''P by(auto elim!: eventually_mono simp add: Q'_def) moreover from Q' G have "\\<^sub>F y in ?Y2. \z. Q (y, z) \ B y z" by(auto simp add: eventually_map_filter_on elim!: eventually_mono) from this[unfolded eq] obtain P'' where P'': "eventually P'' F" and P''Q: "{y. \x. P'' (x, y)} \ {y. \z. ?Q (y, z)}" using F by(fastforce simp add: eventually_map_filter_on) have "eventually (inf P'' ?P) F" using P'' P' by(auto intro: eventually_conj simp add: inf_fun_def) then have "eventually P' F" using P''Q by(auto elim!: eventually_mono simp add: P'_def) ultimately show ?thesis by blast qed show "rel_filter (A OO B) ?X ?Z" proof let ?Y = "\Y. \X Z. eventually X ?X \ eventually Z ?Z \ (\(x, z). X x \ Z z \ (A OO B) x z) \ Y" have Y: "is_filter ?Y" proof show "?Y (\_. True)" by(auto simp add: le_fun_def intro: eventually_True) show "?Y (\x. P x \ Q x)" if "?Y P" "?Y Q" for P Q using that apply clarify apply(intro exI conjI; (elim eventually_rev_mp; fold imp_conjL; intro always_eventually allI; rule imp_refl)?) apply auto done show "?Y Q" if "?Y P" "\x. P x \ Q x" for P Q using that by blast qed define Y where "Y = Abs_filter ?Y" have eventually_Y: "eventually P Y \ ?Y P" for P using eventually_Abs_filter[OF Y, of P] by(simp add: Y_def) show YY: "\\<^sub>F (x, y) in Y. (A OO B) x y" using F G by(auto simp add: eventually_Y eventually_map_filter_on eventually_conj_iff intro!: eventually_True) have "?Y (\(x, z). P x \ (A OO B) x z) \ (\\<^sub>F (x, y) in F. P x \ A x y)" (is "?lhs = ?rhs") for P proof show ?lhs if ?rhs using G F that by(auto 4 3 intro: exI[where x="\_. True"] simp add: eventually_map_filter_on split_def) assume ?lhs then obtain X Z where "\\<^sub>F (x, y) in F. X x \ A x y" and "\\<^sub>F (x, y) in G. Z y \ B x y" and "(\(x, z). X x \ Z z \ (A OO B) x z) \ (\(x, z). P x \ (A OO B) x z)" using F G by(auto simp add: eventually_map_filter_on split_def) from step[OF this(1, 2)] this(3) show ?rhs by(clarsimp elim!: eventually_rev_mp simp add: le_fun_def)(fastforce intro: always_eventually) qed then show "map_filter_on ?AB fst Y = ?X" by(simp add: filter_eq_iff YY eventually_map_filter_on)(simp add: eventually_Y eventually_map_filter_on F G; simp add: split_def) have "?Y (\(x, z). P z \ (A OO B) x z) \ (\\<^sub>F (x, y) in G. P y \ B x y)" (is "?lhs = ?rhs") for P proof show ?lhs if ?rhs using G F that by(auto 4 3 intro: exI[where x="\_. True"] simp add: eventually_map_filter_on split_def) assume ?lhs then obtain X Z where "\\<^sub>F (x, y) in F. X x \ A x y" and "\\<^sub>F (x, y) in G. Z y \ B x y" and "(\(x, z). X x \ Z z \ (A OO B) x z) \ (\(x, z). P z \ (A OO B) x z)" using F G by(auto simp add: eventually_map_filter_on split_def) from step[OF this(1, 2)] this(3) show ?rhs by(clarsimp elim!: eventually_rev_mp simp add: le_fun_def)(fastforce intro: always_eventually) qed then show "map_filter_on ?AB snd Y = ?Z" by(simp add: filter_eq_iff YY eventually_map_filter_on)(simp add: eventually_Y eventually_map_filter_on F G; simp add: split_def) qed qed lemma filtermap_parametric: "((A ===> B) ===> rel_filter A ===> rel_filter B) filtermap filtermap" proof(intro rel_funI; erule rel_filter.cases; hypsubst) fix f g Z assume fg: "(A ===> B) f g" and Z: "\\<^sub>F (x, y) in Z. A x y" have "rel_filter B (map_filter_on {(x, y). A x y} (f \ fst) Z) (map_filter_on {(x, y). A x y} (g \ snd) Z)" (is "rel_filter _ ?F ?G") proof let ?Z = "map_filter_on {(x, y). A x y} (map_prod f g) Z" show "\\<^sub>F (x, y) in ?Z. B x y" using fg Z by(auto simp add: eventually_map_filter_on split_def elim!: eventually_mono rel_funD) have [simp]: "map_prod f g ` {p. A (fst p) (snd p)} \ {p. B (fst p) (snd p)}" using fg by(auto dest: rel_funD) show "map_filter_on {(x, y). B x y} fst ?Z = ?F" "map_filter_on {(x, y). B x y} snd ?Z = ?G" using Z by(auto simp add: map_filter_on_comp split_def) qed thus "rel_filter B (filtermap f (map_filter_on {(x, y). A x y} fst Z)) (filtermap g (map_filter_on {(x, y). A x y} snd Z))" using Z by(simp add: map_filter_on_UNIV[symmetric] map_filter_on_comp) qed lemma rel_filter_Grp: "rel_filter (Grp UNIV f) = Grp UNIV (filtermap f)" proof((intro antisym predicate2I; (elim GrpE; hypsubst)?), rule GrpI[OF _ UNIV_I]) fix F G assume "rel_filter (Grp UNIV f) F G" hence "rel_filter (=) (filtermap f F) (filtermap id G)" by(rule filtermap_parametric[THEN rel_funD, THEN rel_funD, rotated])(simp add: Grp_def rel_fun_def) thus "filtermap f F = G" by(simp add: rel_filter_eq) next fix F :: "'a filter" have "rel_filter (=) F F" by(simp add: rel_filter_eq) hence "rel_filter (Grp UNIV f) (filtermap id F) (filtermap f F)" by(rule filtermap_parametric[THEN rel_funD, THEN rel_funD, rotated])(simp add: Grp_def rel_fun_def) thus "rel_filter (Grp UNIV f) F (filtermap f F)" by simp qed lemma Quotient_filter [quot_map]: "Quotient R Abs Rep T \ Quotient (rel_filter R) (filtermap Abs) (filtermap Rep) (rel_filter T)" unfolding Quotient_alt_def5 rel_filter_eq[symmetric] rel_filter_Grp[symmetric] by(simp add: rel_filter_distr[symmetric] rel_filter_conversep[symmetric] rel_filter_mono) lemma left_total_rel_filter [transfer_rule]: "left_total A \ left_total (rel_filter A)" unfolding left_total_alt_def rel_filter_eq[symmetric] rel_filter_conversep[symmetric] rel_filter_distr by(rule rel_filter_mono) lemma right_total_rel_filter [transfer_rule]: "right_total A \ right_total (rel_filter A)" using left_total_rel_filter[of "A\\"] by(simp add: rel_filter_conversep) lemma bi_total_rel_filter [transfer_rule]: "bi_total A \ bi_total (rel_filter A)" unfolding bi_total_alt_def by(simp add: left_total_rel_filter right_total_rel_filter) lemma left_unique_rel_filter [transfer_rule]: "left_unique A \ left_unique (rel_filter A)" unfolding left_unique_alt_def rel_filter_eq[symmetric] rel_filter_conversep[symmetric] rel_filter_distr by(rule rel_filter_mono) lemma right_unique_rel_filter [transfer_rule]: "right_unique A \ right_unique (rel_filter A)" using left_unique_rel_filter[of "A\\"] by(simp add: rel_filter_conversep) lemma bi_unique_rel_filter [transfer_rule]: "bi_unique A \ bi_unique (rel_filter A)" by(simp add: bi_unique_alt_def left_unique_rel_filter right_unique_rel_filter) lemma eventually_parametric [transfer_rule]: "((A ===> (=)) ===> rel_filter A ===> (=)) eventually eventually" by(auto 4 4 intro!: rel_funI elim!: rel_filter.cases simp add: eventually_map_filter_on dest: rel_funD intro: always_eventually elim!: eventually_rev_mp) lemma frequently_parametric [transfer_rule]: "((A ===> (=)) ===> rel_filter A ===> (=)) frequently frequently" unfolding frequently_def[abs_def] by transfer_prover lemma is_filter_parametric [transfer_rule]: assumes [transfer_rule]: "bi_total A" assumes [transfer_rule]: "bi_unique A" shows "(((A ===> (=)) ===> (=)) ===> (=)) is_filter is_filter" unfolding is_filter_def by transfer_prover lemma top_filter_parametric [transfer_rule]: "rel_filter A top top" if "bi_total A" proof let ?Z = "principal {(x, y). A x y}" show "\\<^sub>F (x, y) in ?Z. A x y" by(simp add: eventually_principal) show "map_filter_on {(x, y). A x y} fst ?Z = top" "map_filter_on {(x, y). A x y} snd ?Z = top" using that by(auto simp add: filter_eq_iff eventually_map_filter_on eventually_principal bi_total_def) qed lemma bot_filter_parametric [transfer_rule]: "rel_filter A bot bot" proof show "\\<^sub>F (x, y) in bot. A x y" by simp show "map_filter_on {(x, y). A x y} fst bot = bot" "map_filter_on {(x, y). A x y} snd bot = bot" by(simp_all add: filter_eq_iff eventually_map_filter_on) qed lemma principal_parametric [transfer_rule]: "(rel_set A ===> rel_filter A) principal principal" proof(rule rel_funI rel_filter.intros)+ fix S S' assume *: "rel_set A S S'" define SS' where "SS' = S \ S' \ {(x, y). A x y}" have SS': "SS' \ {(x, y). A x y}" and [simp]: "S = fst ` SS'" "S' = snd ` SS'" using * by(auto 4 3 dest: rel_setD1 rel_setD2 intro: rev_image_eqI simp add: SS'_def) let ?Z = "principal SS'" show "\\<^sub>F (x, y) in ?Z. A x y" using SS' by(auto simp add: eventually_principal) then show "map_filter_on {(x, y). A x y} fst ?Z = principal S" and "map_filter_on {(x, y). A x y} snd ?Z = principal S'" by(auto simp add: filter_eq_iff eventually_map_filter_on eventually_principal) qed lemma sup_filter_parametric [transfer_rule]: "(rel_filter A ===> rel_filter A ===> rel_filter A) sup sup" proof(intro rel_funI; elim rel_filter.cases; hypsubst) show "rel_filter A (map_filter_on {(x, y). A x y} fst FG \ map_filter_on {(x, y). A x y} fst FG') (map_filter_on {(x, y). A x y} snd FG \ map_filter_on {(x, y). A x y} snd FG')" (is "rel_filter _ (sup ?F ?G) (sup ?F' ?G')") if "\\<^sub>F (x, y) in FG. A x y" "\\<^sub>F (x, y) in FG'. A x y" for FG FG' proof let ?Z = "sup FG FG'" show "\\<^sub>F (x, y) in ?Z. A x y" by(simp add: eventually_sup that) then show "map_filter_on {(x, y). A x y} fst ?Z = sup ?F ?G" and "map_filter_on {(x, y). A x y} snd ?Z = sup ?F' ?G'" by(simp_all add: filter_eq_iff eventually_map_filter_on eventually_sup) qed qed lemma Sup_filter_parametric [transfer_rule]: "(rel_set (rel_filter A) ===> rel_filter A) Sup Sup" proof(rule rel_funI) fix S S' define SS' where "SS' = S \ S' \ {(F, G). rel_filter A F G}" assume "rel_set (rel_filter A) S S'" then have SS': "SS' \ {(F, G). rel_filter A F G}" and [simp]: "S = fst ` SS'" "S' = snd ` SS'" by(auto 4 3 dest: rel_setD1 rel_setD2 intro: rev_image_eqI simp add: SS'_def) from SS' obtain Z where Z: "\F G. (F, G) \ SS' \ (\\<^sub>F (x, y) in Z F G. A x y) \ id F = map_filter_on {(x, y). A x y} fst (Z F G) \ id G = map_filter_on {(x, y). A x y} snd (Z F G)" unfolding rel_filter.simps by atomize_elim((rule choice allI)+; auto) have id: "eventually P F = eventually P (id F)" "eventually Q G = eventually Q (id G)" if "(F, G) \ SS'" for P Q F G by simp_all show "rel_filter A (Sup S) (Sup S')" proof let ?Z = "\(F, G)\SS'. Z F G" show *: "\\<^sub>F (x, y) in ?Z. A x y" using Z by(auto simp add: eventually_Sup) show "map_filter_on {(x, y). A x y} fst ?Z = Sup S" "map_filter_on {(x, y). A x y} snd ?Z = Sup S'" unfolding filter_eq_iff by(auto 4 4 simp add: id eventually_Sup eventually_map_filter_on *[simplified eventually_Sup] simp del: id_apply dest: Z) qed qed context fixes A :: "'a \ 'b \ bool" assumes [transfer_rule]: "bi_unique A" begin lemma le_filter_parametric [transfer_rule]: "(rel_filter A ===> rel_filter A ===> (=)) (\) (\)" unfolding le_filter_def[abs_def] by transfer_prover lemma less_filter_parametric [transfer_rule]: "(rel_filter A ===> rel_filter A ===> (=)) (<) (<)" unfolding less_filter_def[abs_def] by transfer_prover context assumes [transfer_rule]: "bi_total A" begin lemma Inf_filter_parametric [transfer_rule]: "(rel_set (rel_filter A) ===> rel_filter A) Inf Inf" unfolding Inf_filter_def[abs_def] by transfer_prover lemma inf_filter_parametric [transfer_rule]: "(rel_filter A ===> rel_filter A ===> rel_filter A) inf inf" proof(intro rel_funI)+ fix F F' G G' assume [transfer_rule]: "rel_filter A F F'" "rel_filter A G G'" have "rel_filter A (Inf {F, G}) (Inf {F', G'})" by transfer_prover thus "rel_filter A (inf F G) (inf F' G')" by simp qed end end end context includes lifting_syntax begin lemma prod_filter_parametric [transfer_rule]: "(rel_filter R ===> rel_filter S ===> rel_filter (rel_prod R S)) prod_filter prod_filter" proof(intro rel_funI; elim rel_filter.cases; hypsubst) fix F G assume F: "\\<^sub>F (x, y) in F. R x y" and G: "\\<^sub>F (x, y) in G. S x y" show "rel_filter (rel_prod R S) (map_filter_on {(x, y). R x y} fst F \\<^sub>F map_filter_on {(x, y). S x y} fst G) (map_filter_on {(x, y). R x y} snd F \\<^sub>F map_filter_on {(x, y). S x y} snd G)" (is "rel_filter ?RS ?F ?G") proof let ?Z = "filtermap (\((a, b), (a', b')). ((a, a'), (b, b'))) (prod_filter F G)" show *: "\\<^sub>F (x, y) in ?Z. rel_prod R S x y" using F G by(auto simp add: eventually_filtermap split_beta eventually_prod_filter) show "map_filter_on {(x, y). ?RS x y} fst ?Z = ?F" using F G apply(clarsimp simp add: filter_eq_iff eventually_map_filter_on *) apply(simp add: eventually_filtermap split_beta eventually_prod_filter) apply(subst eventually_map_filter_on; simp)+ apply(rule iffI; clarsimp) subgoal for P P' P'' apply(rule exI[where x="\a. \b. P' (a, b) \ R a b"]; rule conjI) subgoal by(fastforce elim: eventually_rev_mp eventually_mono) subgoal by(rule exI[where x="\a. \b. P'' (a, b) \ S a b"])(fastforce elim: eventually_rev_mp eventually_mono) done subgoal by fastforce done show "map_filter_on {(x, y). ?RS x y} snd ?Z = ?G" using F G apply(clarsimp simp add: filter_eq_iff eventually_map_filter_on *) apply(simp add: eventually_filtermap split_beta eventually_prod_filter) apply(subst eventually_map_filter_on; simp)+ apply(rule iffI; clarsimp) subgoal for P P' P'' apply(rule exI[where x="\b. \a. P' (a, b) \ R a b"]; rule conjI) subgoal by(fastforce elim: eventually_rev_mp eventually_mono) subgoal by(rule exI[where x="\b. \a. P'' (a, b) \ S a b"])(fastforce elim: eventually_rev_mp eventually_mono) done subgoal by fastforce done qed qed end text \Code generation for filters\ definition abstract_filter :: "(unit \ 'a filter) \ 'a filter" where [simp]: "abstract_filter f = f ()" code_datatype principal abstract_filter hide_const (open) abstract_filter declare [[code drop: filterlim prod_filter filtermap eventually "inf :: _ filter \ _" "sup :: _ filter \ _" "less_eq :: _ filter \ _" Abs_filter]] declare filterlim_principal [code] declare principal_prod_principal [code] declare filtermap_principal [code] declare filtercomap_principal [code] declare eventually_principal [code] declare inf_principal [code] declare sup_principal [code] declare principal_le_iff [code] lemma Rep_filter_iff_eventually [simp, code]: "Rep_filter F P \ eventually P F" by (simp add: eventually_def) lemma bot_eq_principal_empty [code]: "bot = principal {}" by simp lemma top_eq_principal_UNIV [code]: "top = principal UNIV" by simp instantiation filter :: (equal) equal begin definition equal_filter :: "'a filter \ 'a filter \ bool" where "equal_filter F F' \ F = F'" lemma equal_filter [code]: "HOL.equal (principal A) (principal B) \ A = B" by (simp add: equal_filter_def) instance by standard (simp add: equal_filter_def) end end diff --git a/src/HOL/Library/Countable_Set.thy b/src/HOL/Library/Countable_Set.thy --- a/src/HOL/Library/Countable_Set.thy +++ b/src/HOL/Library/Countable_Set.thy @@ -1,450 +1,451 @@ (* Title: HOL/Library/Countable_Set.thy Author: Johannes Hölzl Author: Andrei Popescu *) section \Countable sets\ theory Countable_Set imports Countable Infinite_Set begin subsection \Predicate for countable sets\ definition countable :: "'a set \ bool" where "countable S \ (\f::'a \ nat. inj_on f S)" lemma countableE: assumes S: "countable S" obtains f :: "'a \ nat" where "inj_on f S" using S by (auto simp: countable_def) lemma countableI: "inj_on (f::'a \ nat) S \ countable S" by (auto simp: countable_def) lemma countableI': "inj_on (f::'a \ 'b::countable) S \ countable S" using comp_inj_on[of f S to_nat] by (auto intro: countableI) lemma countableE_bij: assumes S: "countable S" obtains f :: "nat \ 'a" and C :: "nat set" where "bij_betw f C S" using S by (blast elim: countableE dest: inj_on_imp_bij_betw bij_betw_inv) lemma countableI_bij: "bij_betw f (C::nat set) S \ countable S" by (blast intro: countableI bij_betw_inv_into bij_betw_imp_inj_on) lemma countable_finite: "finite S \ countable S" by (blast dest: finite_imp_inj_to_nat_seg countableI) lemma countableI_bij1: "bij_betw f A B \ countable A \ countable B" by (blast elim: countableE_bij intro: bij_betw_trans countableI_bij) lemma countableI_bij2: "bij_betw f B A \ countable A \ countable B" by (blast elim: countableE_bij intro: bij_betw_trans bij_betw_inv_into countableI_bij) lemma countable_iff_bij[simp]: "bij_betw f A B \ countable A \ countable B" by (blast intro: countableI_bij1 countableI_bij2) lemma countable_subset: "A \ B \ countable B \ countable A" by (auto simp: countable_def intro: subset_inj_on) lemma countableI_type[intro, simp]: "countable (A:: 'a :: countable set)" using countableI[of to_nat A] by auto subsection \Enumerate a countable set\ lemma countableE_infinite: assumes "countable S" "infinite S" obtains e :: "'a \ nat" where "bij_betw e S UNIV" proof - obtain f :: "'a \ nat" where "inj_on f S" using \countable S\ by (rule countableE) then have "bij_betw f S (f`S)" unfolding bij_betw_def by simp moreover from \inj_on f S\ \infinite S\ have inf_fS: "infinite (f`S)" by (auto dest: finite_imageD) then have "bij_betw (the_inv_into UNIV (enumerate (f`S))) (f`S) UNIV" by (intro bij_betw_the_inv_into bij_enumerate) ultimately have "bij_betw (the_inv_into UNIV (enumerate (f`S)) \ f) S UNIV" by (rule bij_betw_trans) then show thesis .. qed lemma countable_infiniteE': assumes "countable A" "infinite A" obtains g where "bij_betw g (UNIV :: nat set) A" using bij_betw_inv[OF countableE_infinite[OF assms]] that by blast lemma countable_enum_cases: assumes "countable S" obtains (finite) f :: "'a \ nat" where "finite S" "bij_betw f S {.. nat" where "infinite S" "bij_betw f S UNIV" using ex_bij_betw_finite_nat[of S] countableE_infinite \countable S\ by (cases "finite S") (auto simp add: atLeast0LessThan) definition to_nat_on :: "'a set \ 'a \ nat" where "to_nat_on S = (SOME f. if finite S then bij_betw f S {..< card S} else bij_betw f S UNIV)" definition from_nat_into :: "'a set \ nat \ 'a" where "from_nat_into S n = (if n \ to_nat_on S ` S then inv_into S (to_nat_on S) n else SOME s. s\S)" lemma to_nat_on_finite: "finite S \ bij_betw (to_nat_on S) S {..< card S}" using ex_bij_betw_finite_nat unfolding to_nat_on_def by (intro someI2_ex[where Q="\f. bij_betw f S {.. infinite S \ bij_betw (to_nat_on S) S UNIV" using countableE_infinite unfolding to_nat_on_def by (intro someI2_ex[where Q="\f. bij_betw f S UNIV"]) auto lemma bij_betw_from_nat_into_finite: "finite S \ bij_betw (from_nat_into S) {..< card S} S" unfolding from_nat_into_def[abs_def] using to_nat_on_finite[of S] apply (subst bij_betw_cong) apply (split if_split) apply (simp add: bij_betw_def) apply (auto cong: bij_betw_cong intro: bij_betw_inv_into to_nat_on_finite) done lemma bij_betw_from_nat_into: "countable S \ infinite S \ bij_betw (from_nat_into S) UNIV S" unfolding from_nat_into_def[abs_def] using to_nat_on_infinite[of S, unfolded bij_betw_def] by (auto cong: bij_betw_cong intro: bij_betw_inv_into to_nat_on_infinite) lemma countable_as_injective_image: assumes "countable A" "infinite A" obtains f :: "nat \ 'a" where "A = range f" "inj f" by (metis bij_betw_def bij_betw_from_nat_into [OF assms]) lemma inj_on_to_nat_on[intro]: "countable A \ inj_on (to_nat_on A) A" using to_nat_on_infinite[of A] to_nat_on_finite[of A] by (cases "finite A") (auto simp: bij_betw_def) lemma to_nat_on_inj[simp]: "countable A \ a \ A \ b \ A \ to_nat_on A a = to_nat_on A b \ a = b" using inj_on_to_nat_on[of A] by (auto dest: inj_onD) lemma from_nat_into_to_nat_on[simp]: "countable A \ a \ A \ from_nat_into A (to_nat_on A a) = a" by (auto simp: from_nat_into_def intro!: inv_into_f_f) lemma subset_range_from_nat_into: "countable A \ A \ range (from_nat_into A)" by (auto intro: from_nat_into_to_nat_on[symmetric]) lemma from_nat_into: "A \ {} \ from_nat_into A n \ A" unfolding from_nat_into_def by (metis equals0I inv_into_into someI_ex) lemma range_from_nat_into_subset: "A \ {} \ range (from_nat_into A) \ A" using from_nat_into[of A] by auto lemma range_from_nat_into[simp]: "A \ {} \ countable A \ range (from_nat_into A) = A" by (metis equalityI range_from_nat_into_subset subset_range_from_nat_into) lemma image_to_nat_on: "countable A \ infinite A \ to_nat_on A ` A = UNIV" using to_nat_on_infinite[of A] by (simp add: bij_betw_def) lemma to_nat_on_surj: "countable A \ infinite A \ \a\A. to_nat_on A a = n" by (metis (no_types) image_iff iso_tuple_UNIV_I image_to_nat_on) lemma to_nat_on_from_nat_into[simp]: "n \ to_nat_on A ` A \ to_nat_on A (from_nat_into A n) = n" by (simp add: f_inv_into_f from_nat_into_def) lemma to_nat_on_from_nat_into_infinite[simp]: "countable A \ infinite A \ to_nat_on A (from_nat_into A n) = n" by (metis image_iff to_nat_on_surj to_nat_on_from_nat_into) lemma from_nat_into_inj: "countable A \ m \ to_nat_on A ` A \ n \ to_nat_on A ` A \ from_nat_into A m = from_nat_into A n \ m = n" by (subst to_nat_on_inj[symmetric, of A]) auto lemma from_nat_into_inj_infinite[simp]: "countable A \ infinite A \ from_nat_into A m = from_nat_into A n \ m = n" using image_to_nat_on[of A] from_nat_into_inj[of A m n] by simp lemma eq_from_nat_into_iff: "countable A \ x \ A \ i \ to_nat_on A ` A \ x = from_nat_into A i \ i = to_nat_on A x" by auto lemma from_nat_into_surj: "countable A \ a \ A \ \n. from_nat_into A n = a" by (rule exI[of _ "to_nat_on A a"]) simp lemma from_nat_into_inject[simp]: "A \ {} \ countable A \ B \ {} \ countable B \ from_nat_into A = from_nat_into B \ A = B" by (metis range_from_nat_into) lemma inj_on_from_nat_into: "inj_on from_nat_into ({A. A \ {} \ countable A})" unfolding inj_on_def by auto subsection \Closure properties of countability\ lemma countable_SIGMA[intro, simp]: "countable I \ (\i. i \ I \ countable (A i)) \ countable (SIGMA i : I. A i)" by (intro countableI'[of "\(i, a). (to_nat_on I i, to_nat_on (A i) a)"]) (auto simp: inj_on_def) lemma countable_image[intro, simp]: assumes "countable A" shows "countable (f`A)" proof - obtain g :: "'a \ nat" where "inj_on g A" using assms by (rule countableE) moreover have "inj_on (inv_into A f) (f`A)" "inv_into A f ` f ` A \ A" by (auto intro: inj_on_inv_into inv_into_into) ultimately show ?thesis by (blast dest: comp_inj_on subset_inj_on intro: countableI) qed lemma countable_image_inj_on: "countable (f ` A) \ inj_on f A \ countable A" by (metis countable_image the_inv_into_onto) lemma countable_image_inj_Int_vimage: "\inj_on f S; countable A\ \ countable (S \ f -` A)" by (meson countable_image_inj_on countable_subset image_subset_iff_subset_vimage inf_le2 inj_on_Int) lemma countable_image_inj_gen: "\inj_on f S; countable A\ \ countable {x \ S. f x \ A}" using countable_image_inj_Int_vimage by (auto simp: vimage_def Collect_conj_eq) lemma countable_image_inj_eq: "inj_on f S \ countable(f ` S) \ countable S" using countable_image_inj_on by blast lemma countable_image_inj: "\countable A; inj f\ \ countable {x. f x \ A}" by (metis (mono_tags, lifting) countable_image_inj_eq countable_subset image_Collect_subsetI inj_on_inverseI the_inv_f_f) lemma countable_UN[intro, simp]: fixes I :: "'i set" and A :: "'i => 'a set" assumes I: "countable I" assumes A: "\i. i \ I \ countable (A i)" shows "countable (\i\I. A i)" proof - have "(\i\I. A i) = snd ` (SIGMA i : I. A i)" by (auto simp: image_iff) then show ?thesis by (simp add: assms) qed lemma countable_Un[intro]: "countable A \ countable B \ countable (A \ B)" by (rule countable_UN[of "{True, False}" "\True \ A | False \ B", simplified]) (simp split: bool.split) lemma countable_Un_iff[simp]: "countable (A \ B) \ countable A \ countable B" by (metis countable_Un countable_subset inf_sup_ord(3,4)) lemma countable_Plus[intro, simp]: "countable A \ countable B \ countable (A <+> B)" by (simp add: Plus_def) lemma countable_empty[intro, simp]: "countable {}" by (blast intro: countable_finite) lemma countable_insert[intro, simp]: "countable A \ countable (insert a A)" using countable_Un[of "{a}" A] by (auto simp: countable_finite) lemma countable_Int1[intro, simp]: "countable A \ countable (A \ B)" by (force intro: countable_subset) lemma countable_Int2[intro, simp]: "countable B \ countable (A \ B)" by (blast intro: countable_subset) lemma countable_INT[intro, simp]: "i \ I \ countable (A i) \ countable (\i\I. A i)" by (blast intro: countable_subset) lemma countable_Diff[intro, simp]: "countable A \ countable (A - B)" by (blast intro: countable_subset) lemma countable_insert_eq [simp]: "countable (insert x A) = countable A" by auto (metis Diff_insert_absorb countable_Diff insert_absorb) lemma countable_vimage: "B \ range f \ countable (f -` B) \ countable B" by (metis Int_absorb2 countable_image image_vimage_eq) lemma surj_countable_vimage: "surj f \ countable (f -` B) \ countable B" by (metis countable_vimage top_greatest) lemma countable_Collect[simp]: "countable A \ countable {a \ A. \ a}" by (metis Collect_conj_eq Int_absorb Int_commute Int_def countable_Int1) lemma countable_Image: assumes "\y. y \ Y \ countable (X `` {y})" assumes "countable Y" shows "countable (X `` Y)" proof - have "countable (X `` (\y\Y. {y}))" unfolding Image_UN by (intro countable_UN assms) then show ?thesis by simp qed lemma countable_relpow: fixes X :: "'a rel" assumes Image_X: "\Y. countable Y \ countable (X `` Y)" assumes Y: "countable Y" shows "countable ((X ^^ i) `` Y)" using Y by (induct i arbitrary: Y) (auto simp: relcomp_Image Image_X) lemma countable_funpow: fixes f :: "'a set \ 'a set" assumes "\A. countable A \ countable (f A)" and "countable A" shows "countable ((f ^^ n) A)" by(induction n)(simp_all add: assms) lemma countable_rtrancl: "(\Y. countable Y \ countable (X `` Y)) \ countable Y \ countable (X\<^sup>* `` Y)" unfolding rtrancl_is_UN_relpow UN_Image by (intro countable_UN countableI_type countable_relpow) lemma countable_lists[intro, simp]: assumes A: "countable A" shows "countable (lists A)" proof - have "countable (lists (range (from_nat_into A)))" by (auto simp: lists_image) with A show ?thesis by (auto dest: subset_range_from_nat_into countable_subset lists_mono) qed lemma Collect_finite_eq_lists: "Collect finite = set ` lists UNIV" using finite_list by auto lemma countable_Collect_finite: "countable (Collect (finite::'a::countable set\bool))" by (simp add: Collect_finite_eq_lists) lemma countable_int: "countable \" unfolding Ints_def by auto lemma countable_rat: "countable \" unfolding Rats_def by auto lemma Collect_finite_subset_eq_lists: "{A. finite A \ A \ T} = set ` lists T" using finite_list by (auto simp: lists_eq_set) lemma countable_Collect_finite_subset: "countable T \ countable {A. finite A \ A \ T}" unfolding Collect_finite_subset_eq_lists by auto lemma countable_Fpow: "countable S \ countable (Fpow S)" using countable_Collect_finite_subset by (force simp add: Fpow_def conj_commute) lemma countable_set_option [simp]: "countable (set_option x)" by (cases x) auto subsection \Misc lemmas\ lemma countable_subset_image: "countable B \ B \ (f ` A) \ (\A'. countable A' \ A' \ A \ (B = f ` A'))" (is "?lhs = ?rhs") proof assume ?lhs show ?rhs by (rule exI [where x="inv_into A f ` B"]) (use \?lhs\ in \auto simp: f_inv_into_f subset_iff image_inv_into_cancel inv_into_into\) next assume ?rhs then show ?lhs by force qed lemma ex_subset_image_inj: "(\T. T \ f ` S \ P T) \ (\T. T \ S \ inj_on f T \ P (f ` T))" by (auto simp: subset_image_inj) lemma all_subset_image_inj: "(\T. T \ f ` S \ P T) \ (\T. T \ S \ inj_on f T \ P(f ` T))" by (metis subset_image_inj) lemma ex_countable_subset_image_inj: "(\T. countable T \ T \ f ` S \ P T) \ (\T. countable T \ T \ S \ inj_on f T \ P (f ` T))" by (metis countable_image_inj_eq subset_image_inj) lemma all_countable_subset_image_inj: "(\T. countable T \ T \ f ` S \ P T) \ (\T. countable T \ T \ S \ inj_on f T \P(f ` T))" by (metis countable_image_inj_eq subset_image_inj) lemma ex_countable_subset_image: "(\T. countable T \ T \ f ` S \ P T) \ (\T. countable T \ T \ S \ P (f ` T))" by (metis countable_subset_image) lemma all_countable_subset_image: "(\T. countable T \ T \ f ` S \ P T) \ (\T. countable T \ T \ S \ P(f ` T))" by (metis countable_subset_image) lemma countable_image_eq: "countable(f ` S) \ (\T. countable T \ T \ S \ f ` S = f ` T)" by (metis countable_image countable_image_inj_eq order_refl subset_image_inj) lemma countable_image_eq_inj: "countable(f ` S) \ (\T. countable T \ T \ S \ f ` S = f ` T \ inj_on f T)" by (metis countable_image_inj_eq order_refl subset_image_inj) lemma infinite_countable_subset': assumes X: "infinite X" shows "\C\X. countable C \ infinite C" proof - - from infinite_countable_subset[OF X] guess f .. + obtain f :: "nat \ 'a" where "inj f" "range f \ X" + using infinite_countable_subset [OF X] by blast then show ?thesis by (intro exI[of _ "range f"]) (auto simp: range_inj_infinite) qed lemma countable_all: assumes S: "countable S" shows "(\s\S. P s) \ (\n::nat. from_nat_into S n \ S \ P (from_nat_into S n))" using S[THEN subset_range_from_nat_into] by auto lemma finite_sequence_to_countable_set: assumes "countable X" obtains F where "\i. F i \ X" "\i. F i \ F (Suc i)" "\i. finite (F i)" "(\i. F i) = X" proof - show thesis apply (rule that[of "\i. if X = {} then {} else from_nat_into X ` {..i}"]) apply (auto simp add: image_iff intro: from_nat_into split: if_splits) using assms from_nat_into_surj by (fastforce cong: image_cong) qed lemma transfer_countable[transfer_rule]: "bi_unique R \ rel_fun (rel_set R) (=) countable countable" by (rule rel_funI, erule (1) bi_unique_rel_set_lemma) (auto dest: countable_image_inj_on) subsection \Uncountable\ abbreviation uncountable where "uncountable A \ \ countable A" lemma uncountable_def: "uncountable A \ A \ {} \ \ (\f::(nat \ 'a). range f = A)" by (auto intro: inj_on_inv_into simp: countable_def) (metis all_not_in_conv inj_on_iff_surj subset_UNIV) lemma uncountable_bij_betw: "bij_betw f A B \ uncountable B \ uncountable A" unfolding bij_betw_def by (metis countable_image) lemma uncountable_infinite: "uncountable A \ infinite A" by (metis countable_finite) lemma uncountable_minus_countable: "uncountable A \ countable B \ uncountable (A - B)" using countable_Un[of B "A - B"] by auto lemma countable_Diff_eq [simp]: "countable (A - {x}) = countable A" by (meson countable_Diff countable_empty countable_insert uncountable_minus_countable) text \Every infinite set can be covered by a pairwise disjoint family of infinite sets. This version doesn't achieve equality, as it only covers a countable subset\ lemma infinite_infinite_partition: assumes "infinite A" obtains C :: "nat \ 'a set" where "pairwise (\i j. disjnt (C i) (C j)) UNIV" "(\i. C i) \ A" "\i. infinite (C i)" proof - obtain f :: "nat\'a" where "range f \ A" "inj f" using assms infinite_countable_subset by blast let ?C = "\i. range (\j. f (prod_encode (i,j)))" show thesis proof show "pairwise (\i j. disjnt (?C i) (?C j)) UNIV" by (auto simp: pairwise_def disjnt_def inj_on_eq_iff [OF \inj f\] inj_on_eq_iff [OF inj_prod_encode, of _ UNIV]) show "(\i. ?C i) \ A" using \range f \ A\ by blast have "infinite (range (\j. f (prod_encode (i, j))))" for i by (rule range_inj_infinite) (meson Pair_inject \inj f\ inj_def prod_encode_eq) then show "\i. infinite (?C i)" using that by auto qed qed end diff --git a/src/HOL/Library/Extended_Nonnegative_Real.thy b/src/HOL/Library/Extended_Nonnegative_Real.thy --- a/src/HOL/Library/Extended_Nonnegative_Real.thy +++ b/src/HOL/Library/Extended_Nonnegative_Real.thy @@ -1,2014 +1,2019 @@ (* Title: HOL/Library/Extended_Nonnegative_Real.thy Author: Johannes Hölzl *) subsection \The type of non-negative extended real numbers\ theory Extended_Nonnegative_Real imports Extended_Real Indicator_Function begin lemma ereal_ineq_diff_add: assumes "b \ (-\::ereal)" "a \ b" shows "a = b + (a-b)" by (metis add.commute assms ereal_eq_minus_iff ereal_minus_le_iff ereal_plus_eq_PInfty) lemma Limsup_const_add: fixes c :: "'a::{complete_linorder, linorder_topology, topological_monoid_add, ordered_ab_semigroup_add}" shows "F \ bot \ Limsup F (\x. c + f x) = c + Limsup F f" by (rule Limsup_compose_continuous_mono) (auto intro!: monoI add_mono continuous_on_add continuous_on_id continuous_on_const) lemma Liminf_const_add: fixes c :: "'a::{complete_linorder, linorder_topology, topological_monoid_add, ordered_ab_semigroup_add}" shows "F \ bot \ Liminf F (\x. c + f x) = c + Liminf F f" by (rule Liminf_compose_continuous_mono) (auto intro!: monoI add_mono continuous_on_add continuous_on_id continuous_on_const) lemma Liminf_add_const: fixes c :: "'a::{complete_linorder, linorder_topology, topological_monoid_add, ordered_ab_semigroup_add}" shows "F \ bot \ Liminf F (\x. f x + c) = Liminf F f + c" by (rule Liminf_compose_continuous_mono) (auto intro!: monoI add_mono continuous_on_add continuous_on_id continuous_on_const) lemma sums_offset: fixes f g :: "nat \ 'a :: {t2_space, topological_comm_monoid_add}" assumes "(\n. f (n + i)) sums l" shows "f sums (l + (\jk. (\nj l + (\jjj=i..j=0..j=i..j\(\n. n + i)`{0..jnjk. (\n l + (\j 'a :: {t2_space, topological_comm_monoid_add}" shows "summable (\j. f (j + i)) \ suminf f = (\j. f (j + i)) + (\jz::real. 0 < z \ z < 1 \ P z) \ eventually P (at_left 1)" by (subst eventually_at_left[of 0]) (auto intro: exI[of _ 0]) lemma mult_eq_1: fixes a b :: "'a :: {ordered_semiring, comm_monoid_mult}" shows "0 \ a \ a \ 1 \ b \ 1 \ a * b = 1 \ (a = 1 \ b = 1)" by (metis mult.left_neutral eq_iff mult.commute mult_right_mono) lemma ereal_add_diff_cancel: fixes a b :: ereal shows "\b\ \ \ \ (a + b) - b = a" by (cases a b rule: ereal2_cases) auto lemma add_top: fixes x :: "'a::{order_top, ordered_comm_monoid_add}" shows "0 \ x \ x + top = top" by (intro top_le add_increasing order_refl) lemma top_add: fixes x :: "'a::{order_top, ordered_comm_monoid_add}" shows "0 \ x \ top + x = top" by (intro top_le add_increasing2 order_refl) lemma le_lfp: "mono f \ x \ lfp f \ f x \ lfp f" by (subst lfp_unfold) (auto dest: monoD) lemma lfp_transfer: assumes \: "sup_continuous \" and f: "sup_continuous f" and mg: "mono g" assumes bot: "\ bot \ lfp g" and eq: "\x. x \ lfp f \ \ (f x) = g (\ x)" shows "\ (lfp f) = lfp g" proof (rule antisym) note mf = sup_continuous_mono[OF f] have f_le_lfp: "(f ^^ i) bot \ lfp f" for i by (induction i) (auto intro: le_lfp mf) have "\ ((f ^^ i) bot) \ lfp g" for i by (induction i) (auto simp: bot eq f_le_lfp intro!: le_lfp mg) then show "\ (lfp f) \ lfp g" unfolding sup_continuous_lfp[OF f] by (subst \[THEN sup_continuousD]) (auto intro!: mono_funpow sup_continuous_mono[OF f] SUP_least) show "lfp g \ \ (lfp f)" by (rule lfp_lowerbound) (simp add: eq[symmetric] lfp_fixpoint[OF mf]) qed lemma sup_continuous_applyD: "sup_continuous f \ sup_continuous (\x. f x h)" using sup_continuous_apply[THEN sup_continuous_compose] . lemma sup_continuous_SUP[order_continuous_intros]: fixes M :: "_ \ _ \ 'a::complete_lattice" assumes M: "\i. i \ I \ sup_continuous (M i)" shows "sup_continuous (SUP i\I. M i)" unfolding sup_continuous_def by (auto simp add: sup_continuousD [OF M] image_comp intro: SUP_commute) lemma sup_continuous_apply_SUP[order_continuous_intros]: fixes M :: "_ \ _ \ 'a::complete_lattice" shows "(\i. i \ I \ sup_continuous (M i)) \ sup_continuous (\x. SUP i\I. M i x)" unfolding SUP_apply[symmetric] by (rule sup_continuous_SUP) lemma sup_continuous_lfp'[order_continuous_intros]: assumes 1: "sup_continuous f" assumes 2: "\g. sup_continuous g \ sup_continuous (f g)" shows "sup_continuous (lfp f)" proof - have "sup_continuous ((f ^^ i) bot)" for i proof (induction i) case (Suc i) then show ?case by (auto intro!: 2) qed (simp add: bot_fun_def sup_continuous_const) then show ?thesis unfolding sup_continuous_lfp[OF 1] by (intro order_continuous_intros) qed lemma sup_continuous_lfp''[order_continuous_intros]: assumes 1: "\s. sup_continuous (f s)" assumes 2: "\g. sup_continuous g \ sup_continuous (\s. f s (g s))" shows "sup_continuous (\x. lfp (f x))" proof - have "sup_continuous (\x. (f x ^^ i) bot)" for i proof (induction i) case (Suc i) then show ?case by (auto intro!: 2) qed (simp add: bot_fun_def sup_continuous_const) then show ?thesis unfolding sup_continuous_lfp[OF 1] by (intro order_continuous_intros) qed lemma mono_INF_fun: "(\x y. mono (F x y)) \ mono (\z x. INF y \ X x. F x y z :: 'a :: complete_lattice)" by (auto intro!: INF_mono[OF bexI] simp: le_fun_def mono_def) lemma continuous_on_cmult_ereal: "\c::ereal\ \ \ \ continuous_on A f \ continuous_on A (\x. c * f x)" using tendsto_cmult_ereal[of c f "f x" "at x within A" for x] by (auto simp: continuous_on_def simp del: tendsto_cmult_ereal) lemma real_of_nat_Sup: assumes "A \ {}" "bdd_above A" shows "of_nat (Sup A) = (SUP a\A. of_nat a :: real)" proof (intro antisym) show "(SUP a\A. of_nat a::real) \ of_nat (Sup A)" using assms by (intro cSUP_least of_nat_mono) (auto intro: cSup_upper) have "Sup A \ A" using assms by (auto simp: Sup_nat_def bdd_above_nat) then show "of_nat (Sup A) \ (SUP a\A. of_nat a::real)" by (intro cSUP_upper bdd_above_image_mono assms) (auto simp: mono_def) qed lemma (in complete_lattice) SUP_sup_const1: "I \ {} \ (SUP i\I. sup c (f i)) = sup c (SUP i\I. f i)" using SUP_sup_distrib[of "\_. c" I f] by simp lemma (in complete_lattice) SUP_sup_const2: "I \ {} \ (SUP i\I. sup (f i) c) = sup (SUP i\I. f i) c" using SUP_sup_distrib[of f I "\_. c"] by simp lemma one_less_of_natD: assumes "(1::'a::linordered_semidom) < of_nat n" shows "1 < n" by (cases n) (use assms in auto) subsection \Defining the extended non-negative reals\ text \Basic definitions and type class setup\ typedef ennreal = "{x :: ereal. 0 \ x}" morphisms enn2ereal e2ennreal' by auto definition "e2ennreal x = e2ennreal' (max 0 x)" lemma enn2ereal_range: "e2ennreal ` {0..} = UNIV" proof - have "\y\0. x = e2ennreal y" for x by (cases x) (auto simp: e2ennreal_def max_absorb2) then show ?thesis by (auto simp: image_iff Bex_def) qed lemma type_definition_ennreal': "type_definition enn2ereal e2ennreal {x. 0 \ x}" using type_definition_ennreal by (auto simp: type_definition_def e2ennreal_def max_absorb2) setup_lifting type_definition_ennreal' declare [[coercion e2ennreal]] instantiation ennreal :: complete_linorder begin lift_definition top_ennreal :: ennreal is top by (rule top_greatest) lift_definition bot_ennreal :: ennreal is 0 by (rule order_refl) lift_definition sup_ennreal :: "ennreal \ ennreal \ ennreal" is sup by (rule le_supI1) lift_definition inf_ennreal :: "ennreal \ ennreal \ ennreal" is inf by (rule le_infI) lift_definition Inf_ennreal :: "ennreal set \ ennreal" is "Inf" by (rule Inf_greatest) lift_definition Sup_ennreal :: "ennreal set \ ennreal" is "sup 0 \ Sup" by auto lift_definition less_eq_ennreal :: "ennreal \ ennreal \ bool" is "(\)" . lift_definition less_ennreal :: "ennreal \ ennreal \ bool" is "(<)" . instance by standard (transfer ; auto simp: Inf_lower Inf_greatest Sup_upper Sup_least le_max_iff_disj max.absorb1)+ end lemma pcr_ennreal_enn2ereal[simp]: "pcr_ennreal (enn2ereal x) x" by (simp add: ennreal.pcr_cr_eq cr_ennreal_def) lemma rel_fun_eq_pcr_ennreal: "rel_fun (=) pcr_ennreal f g \ f = enn2ereal \ g" by (auto simp: rel_fun_def ennreal.pcr_cr_eq cr_ennreal_def) instantiation ennreal :: infinity begin definition infinity_ennreal :: ennreal where [simp]: "\ = (top::ennreal)" instance .. end instantiation ennreal :: "{semiring_1_no_zero_divisors, comm_semiring_1}" begin lift_definition one_ennreal :: ennreal is 1 by simp lift_definition zero_ennreal :: ennreal is 0 by simp lift_definition plus_ennreal :: "ennreal \ ennreal \ ennreal" is "(+)" by simp lift_definition times_ennreal :: "ennreal \ ennreal \ ennreal" is "(*)" by simp instance by standard (transfer; auto simp: field_simps ereal_right_distrib)+ end instantiation ennreal :: minus begin lift_definition minus_ennreal :: "ennreal \ ennreal \ ennreal" is "\a b. max 0 (a - b)" by simp instance .. end instance ennreal :: numeral .. instantiation ennreal :: inverse begin lift_definition inverse_ennreal :: "ennreal \ ennreal" is inverse by (rule inverse_ereal_ge0I) definition divide_ennreal :: "ennreal \ ennreal \ ennreal" where "x div y = x * inverse (y :: ennreal)" instance .. end lemma ennreal_zero_less_one: "0 < (1::ennreal)" \ \TODO: remove\ by transfer auto instance ennreal :: dioid proof (standard; transfer) fix a b :: ereal assume "0 \ a" "0 \ b" then show "(a \ b) = (\c\Collect ((\) 0). b = a + c)" unfolding ereal_ex_split Bex_def by (cases a b rule: ereal2_cases) (auto intro!: exI[of _ "real_of_ereal (b - a)"]) qed instance ennreal :: ordered_comm_semiring by standard (transfer ; auto intro: add_mono mult_mono mult_ac ereal_left_distrib ereal_mult_left_mono)+ instance ennreal :: linordered_nonzero_semiring proof fix a b::ennreal show "a < b \ a + 1 < b + 1" by transfer (simp add: add_right_mono ereal_add_cancel_right less_le) qed (transfer; simp) instance ennreal :: strict_ordered_ab_semigroup_add proof fix a b c d :: ennreal show "a < b \ c < d \ a + c < b + d" by transfer (auto intro!: ereal_add_strict_mono) qed declare [[coercion "of_nat :: nat \ ennreal"]] lemma e2ennreal_neg: "x \ 0 \ e2ennreal x = 0" unfolding zero_ennreal_def e2ennreal_def by (simp add: max_absorb1) lemma e2ennreal_mono: "x \ y \ e2ennreal x \ e2ennreal y" by (cases "0 \ x" "0 \ y" rule: bool.exhaust[case_product bool.exhaust]) (auto simp: e2ennreal_neg less_eq_ennreal.abs_eq eq_onp_def) lemma enn2ereal_nonneg[simp]: "0 \ enn2ereal x" using ennreal.enn2ereal[of x] by simp lemma ereal_ennreal_cases: obtains b where "0 \ a" "a = enn2ereal b" | "a < 0" using e2ennreal'_inverse[of a, symmetric] by (cases "0 \ a") (auto intro: enn2ereal_nonneg) lemma rel_fun_liminf[transfer_rule]: "rel_fun (rel_fun (=) pcr_ennreal) pcr_ennreal liminf liminf" proof - have "rel_fun (rel_fun (=) pcr_ennreal) pcr_ennreal (\x. sup 0 (liminf x)) liminf" unfolding liminf_SUP_INF[abs_def] by (transfer_prover_start, transfer_step+; simp) then show ?thesis apply (subst (asm) (2) rel_fun_def) apply (subst (2) rel_fun_def) apply (auto simp: comp_def max.absorb2 Liminf_bounded rel_fun_eq_pcr_ennreal) done qed lemma rel_fun_limsup[transfer_rule]: "rel_fun (rel_fun (=) pcr_ennreal) pcr_ennreal limsup limsup" proof - have "rel_fun (rel_fun (=) pcr_ennreal) pcr_ennreal (\x. INF n. sup 0 (SUP i\{n..}. x i)) limsup" unfolding limsup_INF_SUP[abs_def] by (transfer_prover_start, transfer_step+; simp) then show ?thesis unfolding limsup_INF_SUP[abs_def] apply (subst (asm) (2) rel_fun_def) apply (subst (2) rel_fun_def) apply (auto simp: comp_def max.absorb2 Sup_upper2 rel_fun_eq_pcr_ennreal) apply (subst (asm) max.absorb2) apply (auto intro: SUP_upper2) done qed lemma sum_enn2ereal[simp]: "(\i. i \ I \ 0 \ f i) \ (\i\I. enn2ereal (f i)) = enn2ereal (sum f I)" by (induction I rule: infinite_finite_induct) (auto simp: sum_nonneg zero_ennreal.rep_eq plus_ennreal.rep_eq) lemma transfer_e2ennreal_sum [transfer_rule]: "rel_fun (rel_fun (=) pcr_ennreal) (rel_fun (=) pcr_ennreal) sum sum" by (auto intro!: rel_funI simp: rel_fun_eq_pcr_ennreal comp_def) lemma enn2ereal_of_nat[simp]: "enn2ereal (of_nat n) = ereal n" by (induction n) (auto simp: zero_ennreal.rep_eq one_ennreal.rep_eq plus_ennreal.rep_eq) lemma enn2ereal_numeral[simp]: "enn2ereal (numeral a) = numeral a" by (metis enn2ereal_of_nat numeral_eq_ereal of_nat_numeral) lemma transfer_numeral[transfer_rule]: "pcr_ennreal (numeral a) (numeral a)" unfolding cr_ennreal_def pcr_ennreal_def by auto subsection \Cancellation simprocs\ lemma ennreal_add_left_cancel: "a + b = a + c \ a = (\::ennreal) \ b = c" unfolding infinity_ennreal_def by transfer (simp add: top_ereal_def ereal_add_cancel_left) lemma ennreal_add_left_cancel_le: "a + b \ a + c \ a = (\::ennreal) \ b \ c" unfolding infinity_ennreal_def by transfer (simp add: ereal_add_le_add_iff top_ereal_def disj_commute) lemma ereal_add_left_cancel_less: fixes a b c :: ereal shows "0 \ a \ 0 \ b \ a + b < a + c \ a \ \ \ b < c" by (cases a b c rule: ereal3_cases) auto lemma ennreal_add_left_cancel_less: "a + b < a + c \ a \ (\::ennreal) \ b < c" unfolding infinity_ennreal_def by transfer (simp add: top_ereal_def ereal_add_left_cancel_less) ML \ structure Cancel_Ennreal_Common = struct (* copied from src/HOL/Tools/nat_numeral_simprocs.ML *) fun find_first_t _ _ [] = raise TERM("find_first_t", []) | find_first_t past u (t::terms) = if u aconv t then (rev past @ terms) else find_first_t (t::past) u terms fun dest_summing (Const (\<^const_name>\Groups.plus\, _) $ t $ u, ts) = dest_summing (t, dest_summing (u, ts)) | dest_summing (t, ts) = t :: ts val mk_sum = Arith_Data.long_mk_sum fun dest_sum t = dest_summing (t, []) val find_first = find_first_t [] val trans_tac = Numeral_Simprocs.trans_tac val norm_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps @{thms ac_simps add_0_left add_0_right}) fun norm_tac ctxt = ALLGOALS (simp_tac (put_simpset norm_ss ctxt)) fun simplify_meta_eq ctxt cancel_th th = Arith_Data.simplify_meta_eq [] ctxt ([th, cancel_th] MRS trans) fun mk_eq (a, b) = HOLogic.mk_Trueprop (HOLogic.mk_eq (a, b)) end structure Eq_Ennreal_Cancel = ExtractCommonTermFun (open Cancel_Ennreal_Common val mk_bal = HOLogic.mk_eq val dest_bal = HOLogic.dest_bin \<^const_name>\HOL.eq\ \<^typ>\ennreal\ fun simp_conv _ _ = SOME @{thm ennreal_add_left_cancel} ) structure Le_Ennreal_Cancel = ExtractCommonTermFun (open Cancel_Ennreal_Common val mk_bal = HOLogic.mk_binrel \<^const_name>\Orderings.less_eq\ val dest_bal = HOLogic.dest_bin \<^const_name>\Orderings.less_eq\ \<^typ>\ennreal\ fun simp_conv _ _ = SOME @{thm ennreal_add_left_cancel_le} ) structure Less_Ennreal_Cancel = ExtractCommonTermFun (open Cancel_Ennreal_Common val mk_bal = HOLogic.mk_binrel \<^const_name>\Orderings.less\ val dest_bal = HOLogic.dest_bin \<^const_name>\Orderings.less\ \<^typ>\ennreal\ fun simp_conv _ _ = SOME @{thm ennreal_add_left_cancel_less} ) \ simproc_setup ennreal_eq_cancel ("(l::ennreal) + m = n" | "(l::ennreal) = m + n") = \fn phi => fn ctxt => fn ct => Eq_Ennreal_Cancel.proc ctxt (Thm.term_of ct)\ simproc_setup ennreal_le_cancel ("(l::ennreal) + m \ n" | "(l::ennreal) \ m + n") = \fn phi => fn ctxt => fn ct => Le_Ennreal_Cancel.proc ctxt (Thm.term_of ct)\ simproc_setup ennreal_less_cancel ("(l::ennreal) + m < n" | "(l::ennreal) < m + n") = \fn phi => fn ctxt => fn ct => Less_Ennreal_Cancel.proc ctxt (Thm.term_of ct)\ subsection \Order with top\ lemma ennreal_zero_less_top[simp]: "0 < (top::ennreal)" by transfer (simp add: top_ereal_def) lemma ennreal_one_less_top[simp]: "1 < (top::ennreal)" by transfer (simp add: top_ereal_def) lemma ennreal_zero_neq_top[simp]: "0 \ (top::ennreal)" by transfer (simp add: top_ereal_def) lemma ennreal_top_neq_zero[simp]: "(top::ennreal) \ 0" by transfer (simp add: top_ereal_def) lemma ennreal_top_neq_one[simp]: "top \ (1::ennreal)" by transfer (simp add: top_ereal_def one_ereal_def flip: ereal_max) lemma ennreal_one_neq_top[simp]: "1 \ (top::ennreal)" by transfer (simp add: top_ereal_def one_ereal_def flip: ereal_max) lemma ennreal_add_less_top[simp]: fixes a b :: ennreal shows "a + b < top \ a < top \ b < top" by transfer (auto simp: top_ereal_def) lemma ennreal_add_eq_top[simp]: fixes a b :: ennreal shows "a + b = top \ a = top \ b = top" by transfer (auto simp: top_ereal_def) lemma ennreal_sum_less_top[simp]: fixes f :: "'a \ ennreal" shows "finite I \ (\i\I. f i) < top \ (\i\I. f i < top)" by (induction I rule: finite_induct) auto lemma ennreal_sum_eq_top[simp]: fixes f :: "'a \ ennreal" shows "finite I \ (\i\I. f i) = top \ (\i\I. f i = top)" by (induction I rule: finite_induct) auto lemma ennreal_mult_eq_top_iff: fixes a b :: ennreal shows "a * b = top \ (a = top \ b \ 0) \ (b = top \ a \ 0)" by transfer (auto simp: top_ereal_def) lemma ennreal_top_eq_mult_iff: fixes a b :: ennreal shows "top = a * b \ (a = top \ b \ 0) \ (b = top \ a \ 0)" using ennreal_mult_eq_top_iff[of a b] by auto lemma ennreal_mult_less_top: fixes a b :: ennreal shows "a * b < top \ (a = 0 \ b = 0 \ (a < top \ b < top))" by transfer (auto simp add: top_ereal_def) lemma top_power_ennreal: "top ^ n = (if n = 0 then 1 else top :: ennreal)" by (induction n) (simp_all add: ennreal_mult_eq_top_iff) lemma ennreal_prod_eq_0[simp]: fixes f :: "'a \ ennreal" shows "(prod f A = 0) = (finite A \ (\i\A. f i = 0))" by (induction A rule: infinite_finite_induct) auto lemma ennreal_prod_eq_top: fixes f :: "'a \ ennreal" shows "(\i\I. f i) = top \ (finite I \ ((\i\I. f i \ 0) \ (\i\I. f i = top)))" by (induction I rule: infinite_finite_induct) (auto simp: ennreal_mult_eq_top_iff) lemma ennreal_top_mult: "top * a = (if a = 0 then 0 else top :: ennreal)" by (simp add: ennreal_mult_eq_top_iff) lemma ennreal_mult_top: "a * top = (if a = 0 then 0 else top :: ennreal)" by (simp add: ennreal_mult_eq_top_iff) lemma enn2ereal_eq_top_iff[simp]: "enn2ereal x = \ \ x = top" by transfer (simp add: top_ereal_def) lemma enn2ereal_top[simp]: "enn2ereal top = \" by transfer (simp add: top_ereal_def) lemma e2ennreal_infty[simp]: "e2ennreal \ = top" by (simp add: top_ennreal.abs_eq top_ereal_def) lemma ennreal_top_minus[simp]: "top - x = (top::ennreal)" by transfer (auto simp: top_ereal_def max_def) lemma minus_top_ennreal: "x - top = (if x = top then top else 0:: ennreal)" by transfer (use ereal_eq_minus_iff top_ereal_def in force) lemma bot_ennreal: "bot = (0::ennreal)" by transfer rule lemma ennreal_of_nat_neq_top[simp]: "of_nat i \ (top::ennreal)" by (induction i) auto lemma numeral_eq_of_nat: "(numeral a::ennreal) = of_nat (numeral a)" by simp lemma of_nat_less_top: "of_nat i < (top::ennreal)" using less_le_trans[of "of_nat i" "of_nat (Suc i)" "top::ennreal"] by simp lemma top_neq_numeral[simp]: "top \ (numeral i::ennreal)" using of_nat_less_top[of "numeral i"] by simp lemma ennreal_numeral_less_top[simp]: "numeral i < (top::ennreal)" using of_nat_less_top[of "numeral i"] by simp lemma ennreal_add_bot[simp]: "bot + x = (x::ennreal)" by transfer simp lemma add_top_right_ennreal [simp]: "x + top = (top :: ennreal)" by (cases x) auto lemma add_top_left_ennreal [simp]: "top + x = (top :: ennreal)" by (cases x) auto lemma ennreal_top_mult_left [simp]: "x \ 0 \ x * top = (top :: ennreal)" by (subst ennreal_mult_eq_top_iff) auto lemma ennreal_top_mult_right [simp]: "x \ 0 \ top * x = (top :: ennreal)" by (subst ennreal_mult_eq_top_iff) auto lemma power_top_ennreal [simp]: "n > 0 \ top ^ n = (top :: ennreal)" by (induction n) auto lemma power_eq_top_ennreal_iff: "x ^ n = top \ x = (top :: ennreal) \ n > 0" by (induction n) (auto simp: ennreal_mult_eq_top_iff) lemma ennreal_mult_le_mult_iff: "c \ 0 \ c \ top \ c * a \ c * b \ a \ (b :: ennreal)" including ennreal.lifting by (transfer, subst ereal_mult_le_mult_iff) (auto simp: top_ereal_def) lemma power_mono_ennreal: "x \ y \ x ^ n \ (y ^ n :: ennreal)" by (induction n) (auto intro!: mult_mono) instance ennreal :: semiring_char_0 proof (standard, safe intro!: linorder_injI) have *: "1 + of_nat k \ (0::ennreal)" for k using add_pos_nonneg[OF zero_less_one, of "of_nat k :: ennreal"] by auto fix x y :: nat assume "x < y" "of_nat x = (of_nat y::ennreal)" then show False by (auto simp add: less_iff_Suc_add *) qed subsection \Arithmetic\ lemma ennreal_minus_zero[simp]: "a - (0::ennreal) = a" by transfer (auto simp: max_def) lemma ennreal_add_diff_cancel_right[simp]: fixes x y z :: ennreal shows "y \ top \ (x + y) - y = x" by transfer (metis ereal_eq_minus_iff max_absorb2 not_MInfty_nonneg top_ereal_def) lemma ennreal_add_diff_cancel_left[simp]: fixes x y z :: ennreal shows "y \ top \ (y + x) - y = x" by (simp add: add.commute) lemma fixes a b :: ennreal shows "a - b = 0 \ a \ b" by transfer (metis ereal_diff_gr0 le_cases max.absorb2 not_less) lemma ennreal_minus_cancel: fixes a b c :: ennreal shows "c \ top \ a \ c \ b \ c \ c - a = c - b \ a = b" apply transfer subgoal for a b c by (cases a b c rule: ereal3_cases) (auto simp: top_ereal_def max_def split: if_splits) done lemma sup_const_add_ennreal: fixes a b c :: "ennreal" shows "sup (c + a) (c + b) = c + sup a b" by transfer (metis add_left_mono le_cases sup.absorb2 sup.orderE) lemma ennreal_diff_add_assoc: fixes a b c :: ennreal shows "a \ b \ c + b - a = c + (b - a)" apply transfer subgoal for a b c by (cases a b c rule: ereal3_cases) (auto simp: field_simps max_absorb2) done lemma mult_divide_eq_ennreal: fixes a b :: ennreal shows "b \ 0 \ b \ top \ (a * b) / b = a" unfolding divide_ennreal_def apply transfer apply (subst mult.assoc) apply (simp add: top_ereal_def flip: divide_ereal_def) done lemma divide_mult_eq: "a \ 0 \ a \ \ \ x * a / (b * a) = x / (b::ennreal)" unfolding divide_ennreal_def infinity_ennreal_def apply transfer subgoal for a b c apply (cases a b c rule: ereal3_cases) apply (auto simp: top_ereal_def) done done lemma ennreal_mult_divide_eq: fixes a b :: ennreal shows "b \ 0 \ b \ top \ (a * b) / b = a" unfolding divide_ennreal_def apply transfer apply (subst mult.assoc) apply (simp add: top_ereal_def flip: divide_ereal_def) done lemma ennreal_add_diff_cancel: fixes a b :: ennreal shows "b \ \ \ (a + b) - b = a" unfolding infinity_ennreal_def by transfer (simp add: max_absorb2 top_ereal_def ereal_add_diff_cancel) lemma ennreal_minus_eq_0: "a - b = 0 \ a \ (b::ennreal)" by transfer (metis ereal_diff_gr0 le_cases max.absorb2 not_less) lemma ennreal_mono_minus_cancel: fixes a b c :: ennreal shows "a - b \ a - c \ a < top \ b \ a \ c \ a \ c \ b" by transfer (auto simp add: max.absorb2 ereal_diff_positive top_ereal_def dest: ereal_mono_minus_cancel) lemma ennreal_mono_minus: fixes a b c :: ennreal shows "c \ b \ a - b \ a - c" by transfer (meson ereal_minus_mono max.mono order_refl) lemma ennreal_minus_pos_iff: fixes a b :: ennreal shows "a < top \ b < top \ 0 < a - b \ b < a" by transfer (use add.left_neutral ereal_minus_le_iff less_irrefl not_less in fastforce) lemma ennreal_inverse_top[simp]: "inverse top = (0::ennreal)" by transfer (simp add: top_ereal_def ereal_inverse_eq_0) lemma ennreal_inverse_zero[simp]: "inverse 0 = (top::ennreal)" by transfer (simp add: top_ereal_def ereal_inverse_eq_0) lemma ennreal_top_divide: "top / (x::ennreal) = (if x = top then 0 else top)" unfolding divide_ennreal_def by transfer (simp add: top_ereal_def ereal_inverse_eq_0 ereal_0_gt_inverse) lemma ennreal_zero_divide[simp]: "0 / (x::ennreal) = 0" by (simp add: divide_ennreal_def) lemma ennreal_divide_zero[simp]: "x / (0::ennreal) = (if x = 0 then 0 else top)" by (simp add: divide_ennreal_def ennreal_mult_top) lemma ennreal_divide_top[simp]: "x / (top::ennreal) = 0" by (simp add: divide_ennreal_def ennreal_top_mult) lemma ennreal_times_divide: "a * (b / c) = a * b / (c::ennreal)" unfolding divide_ennreal_def by transfer (simp add: divide_ereal_def[symmetric] ereal_times_divide_eq) lemma ennreal_zero_less_divide: "0 < a / b \ (0 < a \ b < (top::ennreal))" unfolding divide_ennreal_def by transfer (auto simp: ereal_zero_less_0_iff top_ereal_def ereal_0_gt_inverse) lemma add_divide_distrib_ennreal: "(a + b) / c = a / c + b / (c :: ennreal)" by (simp add: divide_ennreal_def ring_distribs) lemma divide_right_mono_ennreal: fixes a b c :: ennreal shows "a \ b \ a / c \ b / c" unfolding divide_ennreal_def by (intro mult_mono) auto lemma ennreal_mult_strict_right_mono: "(a::ennreal) < c \ 0 < b \ b < top \ a * b < c * b" by transfer (auto intro!: ereal_mult_strict_right_mono) lemma ennreal_indicator_less[simp]: "indicator A x \ (indicator B x::ennreal) \ (x \ A \ x \ B)" by (simp add: indicator_def not_le) lemma ennreal_inverse_positive: "0 < inverse x \ (x::ennreal) \ top" by transfer (simp add: ereal_0_gt_inverse top_ereal_def) lemma ennreal_inverse_mult': "((0 < b \ a < top) \ (0 < a \ b < top)) \ inverse (a * b::ennreal) = inverse a * inverse b" apply transfer subgoal for a b by (cases a b rule: ereal2_cases) (auto simp: top_ereal_def) done lemma ennreal_inverse_mult: "a < top \ b < top \ inverse (a * b::ennreal) = inverse a * inverse b" apply transfer subgoal for a b by (cases a b rule: ereal2_cases) (auto simp: top_ereal_def) done lemma ennreal_inverse_1[simp]: "inverse (1::ennreal) = 1" by transfer simp lemma ennreal_inverse_eq_0_iff[simp]: "inverse (a::ennreal) = 0 \ a = top" by transfer (simp add: ereal_inverse_eq_0 top_ereal_def) lemma ennreal_inverse_eq_top_iff[simp]: "inverse (a::ennreal) = top \ a = 0" by transfer (simp add: top_ereal_def) lemma ennreal_divide_eq_0_iff[simp]: "(a::ennreal) / b = 0 \ (a = 0 \ b = top)" by (simp add: divide_ennreal_def) lemma ennreal_divide_eq_top_iff: "(a::ennreal) / b = top \ ((a \ 0 \ b = 0) \ (a = top \ b \ top))" by (auto simp add: divide_ennreal_def ennreal_mult_eq_top_iff) lemma one_divide_one_divide_ennreal[simp]: "1 / (1 / c) = (c::ennreal)" including ennreal.lifting unfolding divide_ennreal_def by transfer auto lemma ennreal_mult_left_cong: "((a::ennreal) \ 0 \ b = c) \ a * b = a * c" by (cases "a = 0") simp_all lemma ennreal_mult_right_cong: "((a::ennreal) \ 0 \ b = c) \ b * a = c * a" by (cases "a = 0") simp_all lemma ennreal_zero_less_mult_iff: "0 < a * b \ 0 < a \ 0 < (b::ennreal)" by transfer (auto simp add: ereal_zero_less_0_iff le_less) lemma less_diff_eq_ennreal: fixes a b c :: ennreal shows "b < top \ c < top \ a < b - c \ a + c < b" apply transfer subgoal for a b c by (cases a b c rule: ereal3_cases) (auto split: split_max) done lemma diff_add_cancel_ennreal: fixes a b :: ennreal shows "a \ b \ b - a + a = b" unfolding infinity_ennreal_def by transfer (metis (no_types) add.commute ereal_diff_positive ereal_ineq_diff_add max_def not_MInfty_nonneg) lemma ennreal_diff_self[simp]: "a \ top \ a - a = (0::ennreal)" by transfer (simp add: top_ereal_def) lemma ennreal_minus_mono: fixes a b c :: ennreal shows "a \ c \ d \ b \ a - b \ c - d" by transfer (meson ereal_minus_mono max.mono order_refl) lemma ennreal_minus_eq_top[simp]: "a - (b::ennreal) = top \ a = top" by transfer (auto simp: top_ereal_def max.absorb2 ereal_minus_eq_PInfty_iff split: split_max) lemma ennreal_divide_self[simp]: "a \ 0 \ a < top \ a / a = (1::ennreal)" unfolding divide_ennreal_def apply transfer subgoal for a by (cases a) (auto simp: top_ereal_def) done subsection \Coercion from \<^typ>\real\ to \<^typ>\ennreal\\ lift_definition ennreal :: "real \ ennreal" is "sup 0 \ ereal" by simp declare [[coercion ennreal]] lemma ennreal_cong: "x = y \ ennreal x = ennreal y" by simp lemma ennreal_cases[cases type: ennreal]: fixes x :: ennreal obtains (real) r :: real where "0 \ r" "x = ennreal r" | (top) "x = top" apply transfer subgoal for x thesis by (cases x) (auto simp: max.absorb2 top_ereal_def) done lemmas ennreal2_cases = ennreal_cases[case_product ennreal_cases] lemmas ennreal3_cases = ennreal_cases[case_product ennreal2_cases] lemma ennreal_neq_top[simp]: "ennreal r \ top" by transfer (simp add: top_ereal_def zero_ereal_def flip: ereal_max) lemma top_neq_ennreal[simp]: "top \ ennreal r" using ennreal_neq_top[of r] by (auto simp del: ennreal_neq_top) lemma ennreal_less_top[simp]: "ennreal x < top" by transfer (simp add: top_ereal_def max_def) lemma ennreal_neg: "x \ 0 \ ennreal x = 0" by transfer (simp add: max.absorb1) lemma ennreal_inj[simp]: "0 \ a \ 0 \ b \ ennreal a = ennreal b \ a = b" by (transfer fixing: a b) (auto simp: max_absorb2) lemma ennreal_le_iff[simp]: "0 \ y \ ennreal x \ ennreal y \ x \ y" by (auto simp: ennreal_def zero_ereal_def less_eq_ennreal.abs_eq eq_onp_def split: split_max) lemma le_ennreal_iff: "0 \ r \ x \ ennreal r \ (\q\0. x = ennreal q \ q \ r)" by (cases x) (auto simp: top_unique) lemma ennreal_less_iff: "0 \ r \ ennreal r < ennreal q \ r < q" unfolding not_le[symmetric] by auto lemma ennreal_eq_zero_iff[simp]: "0 \ x \ ennreal x = 0 \ x = 0" by transfer (auto simp: max_absorb2) lemma ennreal_less_zero_iff[simp]: "0 < ennreal x \ 0 < x" by transfer (auto simp: max_def) lemma ennreal_lessI: "0 < q \ r < q \ ennreal r < ennreal q" by (cases "0 \ r") (auto simp: ennreal_less_iff ennreal_neg) lemma ennreal_leI: "x \ y \ ennreal x \ ennreal y" by (cases "0 \ y") (auto simp: ennreal_neg) lemma enn2ereal_ennreal[simp]: "0 \ x \ enn2ereal (ennreal x) = x" by transfer (simp add: max_absorb2) lemma e2ennreal_enn2ereal[simp]: "e2ennreal (enn2ereal x) = x" by (simp add: e2ennreal_def max_absorb2 ennreal.enn2ereal_inverse) lemma enn2ereal_e2ennreal: "x \ 0 \ enn2ereal (e2ennreal x) = x" by (metis e2ennreal_enn2ereal ereal_ennreal_cases not_le) lemma e2ennreal_ereal [simp]: "e2ennreal (ereal x) = ennreal x" by (metis e2ennreal_def enn2ereal_inverse ennreal.rep_eq sup_ereal_def) lemma ennreal_0[simp]: "ennreal 0 = 0" by (simp add: ennreal_def max.absorb1 zero_ennreal.abs_eq) lemma ennreal_1[simp]: "ennreal 1 = 1" by transfer (simp add: max_absorb2) lemma ennreal_eq_0_iff: "ennreal x = 0 \ x \ 0" by (cases "0 \ x") (auto simp: ennreal_neg) lemma ennreal_le_iff2: "ennreal x \ ennreal y \ ((0 \ y \ x \ y) \ (x \ 0 \ y \ 0))" by (cases "0 \ y") (auto simp: ennreal_eq_0_iff ennreal_neg) lemma ennreal_eq_1[simp]: "ennreal x = 1 \ x = 1" by (cases "0 \ x") (auto simp: ennreal_neg simp flip: ennreal_1) lemma ennreal_le_1[simp]: "ennreal x \ 1 \ x \ 1" by (cases "0 \ x") (auto simp: ennreal_neg simp flip: ennreal_1) lemma ennreal_ge_1[simp]: "ennreal x \ 1 \ x \ 1" by (cases "0 \ x") (auto simp: ennreal_neg simp flip: ennreal_1) lemma one_less_ennreal[simp]: "1 < ennreal x \ 1 < x" by transfer (auto simp: max.absorb2 less_max_iff_disj) lemma ennreal_plus[simp]: "0 \ a \ 0 \ b \ ennreal (a + b) = ennreal a + ennreal b" by (transfer fixing: a b) (auto simp: max_absorb2) lemma add_mono_ennreal: "x < ennreal y \ x' < ennreal y' \ x + x' < ennreal (y + y')" by (metis (full_types) add_strict_mono ennreal_less_zero_iff ennreal_plus less_le not_less zero_le) lemma sum_ennreal[simp]: "(\i. i \ I \ 0 \ f i) \ (\i\I. ennreal (f i)) = ennreal (sum f I)" by (induction I rule: infinite_finite_induct) (auto simp: sum_nonneg) lemma sum_list_ennreal[simp]: assumes "\x. x \ set xs \ f x \ 0" shows "sum_list (map (\x. ennreal (f x)) xs) = ennreal (sum_list (map f xs))" using assms proof (induction xs) case (Cons x xs) from Cons have "(\x\x # xs. ennreal (f x)) = ennreal (f x) + ennreal (sum_list (map f xs))" by simp also from Cons.prems have "\ = ennreal (f x + sum_list (map f xs))" by (intro ennreal_plus [symmetric] sum_list_nonneg) auto finally show ?case by simp qed simp_all lemma ennreal_of_nat_eq_real_of_nat: "of_nat i = ennreal (of_nat i)" by (induction i) simp_all lemma of_nat_le_ennreal_iff[simp]: "0 \ r \ of_nat i \ ennreal r \ of_nat i \ r" by (simp add: ennreal_of_nat_eq_real_of_nat) lemma ennreal_le_of_nat_iff[simp]: "ennreal r \ of_nat i \ r \ of_nat i" by (simp add: ennreal_of_nat_eq_real_of_nat) lemma ennreal_indicator: "ennreal (indicator A x) = indicator A x" by (auto split: split_indicator) lemma ennreal_numeral[simp]: "ennreal (numeral n) = numeral n" using ennreal_of_nat_eq_real_of_nat[of "numeral n"] by simp lemma ennreal_less_numeral_iff [simp]: "ennreal n < numeral w \ n < numeral w" by (metis ennreal_less_iff ennreal_numeral less_le not_less zero_less_numeral) lemma numeral_less_ennreal_iff [simp]: "numeral w < ennreal n \ numeral w < n" using ennreal_less_iff zero_le_numeral by fastforce lemma numeral_le_ennreal_iff [simp]: "numeral n \ ennreal m \ numeral n \ m" by (metis not_le ennreal_less_numeral_iff) lemma min_ennreal: "0 \ x \ 0 \ y \ min (ennreal x) (ennreal y) = ennreal (min x y)" by (auto split: split_min) lemma ennreal_half[simp]: "ennreal (1/2) = inverse 2" by transfer (simp add: max.absorb2) lemma ennreal_minus: "0 \ q \ ennreal r - ennreal q = ennreal (r - q)" by transfer (simp add: max.absorb2 zero_ereal_def flip: ereal_max) lemma ennreal_minus_top[simp]: "ennreal a - top = 0" by (simp add: minus_top_ennreal) lemma e2eenreal_enn2ereal_diff [simp]: "e2ennreal(enn2ereal x - enn2ereal y) = x - y" for x y by (cases x, cases y, auto simp add: ennreal_minus e2ennreal_neg) lemma ennreal_mult: "0 \ a \ 0 \ b \ ennreal (a * b) = ennreal a * ennreal b" by transfer (simp add: max_absorb2) lemma ennreal_mult': "0 \ a \ ennreal (a * b) = ennreal a * ennreal b" by (cases "0 \ b") (auto simp: ennreal_mult ennreal_neg mult_nonneg_nonpos) lemma indicator_mult_ennreal: "indicator A x * ennreal r = ennreal (indicator A x * r)" by (simp split: split_indicator) lemma ennreal_mult'': "0 \ b \ ennreal (a * b) = ennreal a * ennreal b" by (cases "0 \ a") (auto simp: ennreal_mult ennreal_neg mult_nonpos_nonneg) lemma numeral_mult_ennreal: "0 \ x \ numeral b * ennreal x = ennreal (numeral b * x)" by (simp add: ennreal_mult) lemma ennreal_power: "0 \ r \ ennreal r ^ n = ennreal (r ^ n)" by (induction n) (auto simp: ennreal_mult) lemma power_eq_top_ennreal: "x ^ n = top \ (n \ 0 \ (x::ennreal) = top)" by (cases x rule: ennreal_cases) (auto simp: ennreal_power top_power_ennreal) lemma inverse_ennreal: "0 < r \ inverse (ennreal r) = ennreal (inverse r)" by transfer (simp add: max.absorb2) lemma divide_ennreal: "0 \ r \ 0 < q \ ennreal r / ennreal q = ennreal (r / q)" by (simp add: divide_ennreal_def inverse_ennreal ennreal_mult[symmetric] inverse_eq_divide) lemma ennreal_inverse_power: "inverse (x ^ n :: ennreal) = inverse x ^ n" proof (cases x rule: ennreal_cases) case top with power_eq_top_ennreal[of x n] show ?thesis by (cases "n = 0") auto next case (real r) then show ?thesis proof cases assume "x = 0" then show ?thesis using power_eq_top_ennreal[of top "n - 1"] by (cases n) (auto simp: ennreal_top_mult) next assume "x \ 0" with real have "0 < r" by auto with real show ?thesis by (induction n) (auto simp add: ennreal_power ennreal_mult[symmetric] inverse_ennreal) qed qed lemma power_divide_distrib_ennreal [algebra_simps]: "(x / y) ^ n = x ^ n / (y ^ n :: ennreal)" by (simp add: divide_ennreal_def algebra_simps ennreal_inverse_power) lemma ennreal_divide_numeral: "0 \ x \ ennreal x / numeral b = ennreal (x / numeral b)" by (subst divide_ennreal[symmetric]) auto lemma prod_ennreal: "(\i. i \ A \ 0 \ f i) \ (\i\A. ennreal (f i)) = ennreal (prod f A)" by (induction A rule: infinite_finite_induct) (auto simp: ennreal_mult prod_nonneg) lemma prod_mono_ennreal: assumes "\x. x \ A \ f x \ (g x :: ennreal)" shows "prod f A \ prod g A" using assms by (induction A rule: infinite_finite_induct) (auto intro!: mult_mono) lemma mult_right_ennreal_cancel: "a * ennreal c = b * ennreal c \ (a = b \ c \ 0)" proof (cases "0 \ c") case True then show ?thesis by (metis ennreal_eq_0_iff ennreal_mult_right_cong ennreal_neq_top mult_divide_eq_ennreal) qed (use ennreal_neg in auto) lemma ennreal_le_epsilon: "(\e::real. y < top \ 0 < e \ x \ y + ennreal e) \ x \ y" apply (cases y rule: ennreal_cases) apply (cases x rule: ennreal_cases) apply (auto simp flip: ennreal_plus simp add: top_unique intro: zero_less_one field_le_epsilon) done lemma ennreal_rat_dense: fixes x y :: ennreal shows "x < y \ \r::rat. x < real_of_rat r \ real_of_rat r < y" proof transfer fix x y :: ereal assume xy: "0 \ x" "0 \ y" "x < y" moreover from ereal_dense3[OF \x < y\] obtain r where r: "x < ereal (real_of_rat r)" "ereal (real_of_rat r) < y" by auto then have "0 \ r" using le_less_trans[OF \0 \ x\ \x < ereal (real_of_rat r)\] by auto with r show "\r. x < (sup 0 \ ereal) (real_of_rat r) \ (sup 0 \ ereal) (real_of_rat r) < y" by (intro exI[of _ r]) (auto simp: max_absorb2) qed lemma ennreal_Ex_less_of_nat: "(x::ennreal) < top \ \n. x < of_nat n" by (cases x rule: ennreal_cases) (auto simp: ennreal_of_nat_eq_real_of_nat ennreal_less_iff reals_Archimedean2) subsection \Coercion from \<^typ>\ennreal\ to \<^typ>\real\\ definition "enn2real x = real_of_ereal (enn2ereal x)" lemma enn2real_nonneg[simp]: "0 \ enn2real x" by (auto simp: enn2real_def intro!: real_of_ereal_pos enn2ereal_nonneg) lemma enn2real_mono: "a \ b \ b < top \ enn2real a \ enn2real b" by (auto simp add: enn2real_def less_eq_ennreal.rep_eq intro!: real_of_ereal_positive_mono enn2ereal_nonneg) lemma enn2real_of_nat[simp]: "enn2real (of_nat n) = n" by (auto simp: enn2real_def) lemma enn2real_ennreal[simp]: "0 \ r \ enn2real (ennreal r) = r" by (simp add: enn2real_def) lemma ennreal_enn2real[simp]: "r < top \ ennreal (enn2real r) = r" by (cases r rule: ennreal_cases) auto lemma real_of_ereal_enn2ereal[simp]: "real_of_ereal (enn2ereal x) = enn2real x" by (simp add: enn2real_def) lemma enn2real_top[simp]: "enn2real top = 0" unfolding enn2real_def top_ennreal.rep_eq top_ereal_def by simp lemma enn2real_0[simp]: "enn2real 0 = 0" unfolding enn2real_def zero_ennreal.rep_eq by simp lemma enn2real_1[simp]: "enn2real 1 = 1" unfolding enn2real_def one_ennreal.rep_eq by simp lemma enn2real_numeral[simp]: "enn2real (numeral n) = (numeral n)" unfolding enn2real_def by simp lemma enn2real_mult: "enn2real (a * b) = enn2real a * enn2real b" unfolding enn2real_def by (simp del: real_of_ereal_enn2ereal add: times_ennreal.rep_eq) lemma enn2real_leI: "0 \ B \ x \ ennreal B \ enn2real x \ B" by (cases x rule: ennreal_cases) (auto simp: top_unique) lemma enn2real_positive_iff: "0 < enn2real x \ (0 < x \ x < top)" by (cases x rule: ennreal_cases) auto lemma enn2real_eq_posreal_iff[simp]: "c > 0 \ enn2real x = c \ x = c" by (cases x) auto lemma ennreal_enn2real_if: "ennreal (enn2real r) = (if r = top then 0 else r)" by(auto intro!: ennreal_enn2real simp add: less_top) subsection \Coercion from \<^typ>\enat\ to \<^typ>\ennreal\\ definition ennreal_of_enat :: "enat \ ennreal" where "ennreal_of_enat n = (case n of \ \ top | enat n \ of_nat n)" declare [[coercion ennreal_of_enat]] declare [[coercion "of_nat :: nat \ ennreal"]] lemma ennreal_of_enat_infty[simp]: "ennreal_of_enat \ = \" by (simp add: ennreal_of_enat_def) lemma ennreal_of_enat_enat[simp]: "ennreal_of_enat (enat n) = of_nat n" by (simp add: ennreal_of_enat_def) lemma ennreal_of_enat_0[simp]: "ennreal_of_enat 0 = 0" using ennreal_of_enat_enat[of 0] unfolding enat_0 by simp lemma ennreal_of_enat_1[simp]: "ennreal_of_enat 1 = 1" using ennreal_of_enat_enat[of 1] unfolding enat_1 by simp lemma ennreal_top_neq_of_nat[simp]: "(top::ennreal) \ of_nat i" using ennreal_of_nat_neq_top[of i] by metis lemma ennreal_of_enat_inj[simp]: "ennreal_of_enat i = ennreal_of_enat j \ i = j" by (cases i j rule: enat.exhaust[case_product enat.exhaust]) auto lemma ennreal_of_enat_le_iff[simp]: "ennreal_of_enat m \ ennreal_of_enat n \ m \ n" by (auto simp: ennreal_of_enat_def top_unique split: enat.split) lemma of_nat_less_ennreal_of_nat[simp]: "of_nat n \ ennreal_of_enat x \ of_nat n \ x" by (cases x) (auto simp: of_nat_eq_enat) lemma ennreal_of_enat_Sup: "ennreal_of_enat (Sup X) = (SUP x\X. ennreal_of_enat x)" proof - have "ennreal_of_enat (Sup X) \ (SUP x \ X. ennreal_of_enat x)" unfolding Sup_enat_def proof (clarsimp, intro conjI impI) fix x assume "finite X" "X \ {}" then show "ennreal_of_enat (Max X) \ (SUP x \ X. ennreal_of_enat x)" by (intro SUP_upper Max_in) next assume "infinite X" "X \ {}" have "\y\X. r < ennreal_of_enat y" if r: "r < top" for r proof - - from ennreal_Ex_less_of_nat[OF r] guess n .. note n = this + obtain n where n: "r < of_nat n" + using ennreal_Ex_less_of_nat[OF r] .. have "\ (X \ enat ` {.. n})" using \infinite X\ by (auto dest: finite_subset) then obtain x where x: "x \ X" "x \ enat ` {..n}" by blast then have "of_nat n \ x" by (cases x) (auto simp: of_nat_eq_enat) with x show ?thesis by (auto intro!: bexI[of _ x] less_le_trans[OF n]) qed then have "(SUP x \ X. ennreal_of_enat x) = top" by simp then show "top \ (SUP x \ X. ennreal_of_enat x)" unfolding top_unique by simp qed then show ?thesis by (auto intro!: antisym Sup_least intro: Sup_upper) qed lemma ennreal_of_enat_eSuc[simp]: "ennreal_of_enat (eSuc x) = 1 + ennreal_of_enat x" by (cases x) (auto simp: eSuc_enat) subsection \Topology on \<^typ>\ennreal\\ lemma enn2ereal_Iio: "enn2ereal -` {.. a then {..< e2ennreal a} else {})" using enn2ereal_nonneg by (cases a rule: ereal_ennreal_cases) (auto simp add: vimage_def set_eq_iff ennreal.enn2ereal_inverse less_ennreal.rep_eq e2ennreal_def max_absorb2 simp del: enn2ereal_nonneg intro: le_less_trans less_imp_le) lemma enn2ereal_Ioi: "enn2ereal -` {a <..} = (if 0 \ a then {e2ennreal a <..} else UNIV)" by (cases a rule: ereal_ennreal_cases) (auto simp add: vimage_def set_eq_iff ennreal.enn2ereal_inverse less_ennreal.rep_eq e2ennreal_def max_absorb2 intro: less_le_trans) instantiation ennreal :: linear_continuum_topology begin definition open_ennreal :: "ennreal set \ bool" where "(open :: ennreal set \ bool) = generate_topology (range lessThan \ range greaterThan)" instance proof show "\a b::ennreal. a \ b" using zero_neq_one by (intro exI) show "\x y::ennreal. x < y \ \z>x. z < y" proof transfer - fix x y :: ereal assume "0 \ x" and *: "x < y" - moreover from dense[OF *] guess z .. - ultimately show "\z\Collect ((\) 0). x < z \ z < y" + fix x y :: ereal + assume *: "0 \ x" + assume "x < y" + from dense[OF this] obtain z where "x < z \ z < y" .. + with * show "\z\Collect ((\) 0). x < z \ z < y" by (intro bexI[of _ z]) auto qed qed (rule open_ennreal_def) end lemma continuous_on_e2ennreal: "continuous_on A e2ennreal" proof (rule continuous_on_subset) show "continuous_on ({0..} \ {..0}) e2ennreal" proof (rule continuous_on_closed_Un) show "continuous_on {0 ..} e2ennreal" by (rule continuous_onI_mono) (auto simp add: less_eq_ennreal.abs_eq eq_onp_def enn2ereal_range) show "continuous_on {.. 0} e2ennreal" by (subst continuous_on_cong[OF refl, of _ _ "\_. 0"]) (auto simp add: e2ennreal_neg continuous_on_const) qed auto show "A \ {0..} \ {..0::ereal}" by auto qed lemma continuous_at_e2ennreal: "continuous (at x within A) e2ennreal" by (rule continuous_on_imp_continuous_within[OF continuous_on_e2ennreal, of _ UNIV]) auto lemma continuous_on_enn2ereal: "continuous_on UNIV enn2ereal" by (rule continuous_on_generate_topology[OF open_generated_order]) (auto simp add: enn2ereal_Iio enn2ereal_Ioi) lemma continuous_at_enn2ereal: "continuous (at x within A) enn2ereal" by (rule continuous_on_imp_continuous_within[OF continuous_on_enn2ereal]) auto lemma sup_continuous_e2ennreal[order_continuous_intros]: assumes f: "sup_continuous f" shows "sup_continuous (\x. e2ennreal (f x))" proof (rule sup_continuous_compose[OF _ f]) show "sup_continuous e2ennreal" by (simp add: continuous_at_e2ennreal continuous_at_left_imp_sup_continuous e2ennreal_mono mono_def) qed lemma sup_continuous_enn2ereal[order_continuous_intros]: assumes f: "sup_continuous f" shows "sup_continuous (\x. enn2ereal (f x))" proof (rule sup_continuous_compose[OF _ f]) show "sup_continuous enn2ereal" by (simp add: continuous_at_enn2ereal continuous_at_left_imp_sup_continuous less_eq_ennreal.rep_eq mono_def) qed lemma sup_continuous_mult_left_ennreal': fixes c :: "ennreal" shows "sup_continuous (\x. c * x)" unfolding sup_continuous_def by transfer (auto simp: SUP_ereal_mult_left max.absorb2 SUP_upper2) lemma sup_continuous_mult_left_ennreal[order_continuous_intros]: "sup_continuous f \ sup_continuous (\x. c * f x :: ennreal)" by (rule sup_continuous_compose[OF sup_continuous_mult_left_ennreal']) lemma sup_continuous_mult_right_ennreal[order_continuous_intros]: "sup_continuous f \ sup_continuous (\x. f x * c :: ennreal)" using sup_continuous_mult_left_ennreal[of f c] by (simp add: mult.commute) lemma sup_continuous_divide_ennreal[order_continuous_intros]: fixes f g :: "'a::complete_lattice \ ennreal" shows "sup_continuous f \ sup_continuous (\x. f x / c)" unfolding divide_ennreal_def by (rule sup_continuous_mult_right_ennreal) lemma transfer_enn2ereal_continuous_on [transfer_rule]: "rel_fun (=) (rel_fun (rel_fun (=) pcr_ennreal) (=)) continuous_on continuous_on" proof - have "continuous_on A f" if "continuous_on A (\x. enn2ereal (f x))" for A and f :: "'a \ ennreal" using continuous_on_compose2[OF continuous_on_e2ennreal[of "{0..}"] that] by (auto simp: ennreal.enn2ereal_inverse subset_eq e2ennreal_def max_absorb2) moreover have "continuous_on A (\x. enn2ereal (f x))" if "continuous_on A f" for A and f :: "'a \ ennreal" using continuous_on_compose2[OF continuous_on_enn2ereal that] by auto ultimately show ?thesis by (auto simp add: rel_fun_def ennreal.pcr_cr_eq cr_ennreal_def) qed lemma transfer_sup_continuous[transfer_rule]: "(rel_fun (rel_fun (=) pcr_ennreal) (=)) sup_continuous sup_continuous" proof (safe intro!: rel_funI dest!: rel_fun_eq_pcr_ennreal[THEN iffD1]) show "sup_continuous (enn2ereal \ f) \ sup_continuous f" for f :: "'a \ _" using sup_continuous_e2ennreal[of "enn2ereal \ f"] by simp show "sup_continuous f \ sup_continuous (enn2ereal \ f)" for f :: "'a \ _" using sup_continuous_enn2ereal[of f] by (simp add: comp_def) qed lemma continuous_on_ennreal[tendsto_intros]: "continuous_on A f \ continuous_on A (\x. ennreal (f x))" by transfer (auto intro!: continuous_on_max continuous_on_const continuous_on_ereal) lemma tendsto_ennrealD: assumes lim: "((\x. ennreal (f x)) \ ennreal x) F" assumes *: "\\<^sub>F x in F. 0 \ f x" and x: "0 \ x" shows "(f \ x) F" proof - have "((\x. enn2ereal (ennreal (f x))) \ enn2ereal (ennreal x)) F \ (f \ enn2ereal (ennreal x)) F" using "*" eventually_mono by (intro tendsto_cong) fastforce then show ?thesis using assms(1) continuous_at_enn2ereal isCont_tendsto_compose x by fastforce qed lemma tendsto_ennreal_iff [simp]: \((\x. ennreal (f x)) \ ennreal x) F \ (f \ x) F\ (is \?P \ ?Q\) if \\\<^sub>F x in F. 0 \ f x\ \0 \ x\ proof assume \?P\ then show \?Q\ using that by (rule tendsto_ennrealD) next assume \?Q\ have \continuous_on UNIV ereal\ using continuous_on_ereal [of _ id] by simp then have \continuous_on UNIV (e2ennreal \ ereal)\ by (rule continuous_on_compose) (simp_all add: continuous_on_e2ennreal) then have \((\x. (e2ennreal \ ereal) (f x)) \ (e2ennreal \ ereal) x) F\ using \?Q\ by (rule continuous_on_tendsto_compose) simp_all then show \?P\ by (simp flip: e2ennreal_ereal) qed lemma tendsto_enn2ereal_iff[simp]: "((\i. enn2ereal (f i)) \ enn2ereal x) F \ (f \ x) F" using continuous_on_enn2ereal[THEN continuous_on_tendsto_compose, of f x F] continuous_on_e2ennreal[THEN continuous_on_tendsto_compose, of "\x. enn2ereal (f x)" "enn2ereal x" F UNIV] by auto lemma ennreal_tendsto_0_iff: "(\n. f n \ 0) \ ((\n. ennreal (f n)) \ 0) \ (f \ 0)" by (metis (mono_tags) ennreal_0 eventuallyI order_refl tendsto_ennreal_iff) lemma continuous_on_add_ennreal: fixes f g :: "'a::topological_space \ ennreal" shows "continuous_on A f \ continuous_on A g \ continuous_on A (\x. f x + g x)" by (transfer fixing: A) (auto intro!: tendsto_add_ereal_nonneg simp: continuous_on_def) lemma continuous_on_inverse_ennreal[continuous_intros]: fixes f :: "'a::topological_space \ ennreal" shows "continuous_on A f \ continuous_on A (\x. inverse (f x))" proof (transfer fixing: A) show "pred_fun top ((\) 0) f \ continuous_on A (\x. inverse (f x))" if "continuous_on A f" for f :: "'a \ ereal" using continuous_on_compose2[OF continuous_on_inverse_ereal that] by (auto simp: subset_eq) qed instance ennreal :: topological_comm_monoid_add proof show "((\x. fst x + snd x) \ a + b) (nhds a \\<^sub>F nhds b)" for a b :: ennreal using continuous_on_add_ennreal[of UNIV fst snd] using tendsto_at_iff_tendsto_nhds[symmetric, of "\x::(ennreal \ ennreal). fst x + snd x"] by (auto simp: continuous_on_eq_continuous_at) (simp add: isCont_def nhds_prod[symmetric]) qed lemma sup_continuous_add_ennreal[order_continuous_intros]: fixes f g :: "'a::complete_lattice \ ennreal" shows "sup_continuous f \ sup_continuous g \ sup_continuous (\x. f x + g x)" by transfer (auto intro!: sup_continuous_add) lemma ennreal_suminf_lessD: "(\i. f i :: ennreal) < x \ f i < x" using le_less_trans[OF sum_le_suminf[OF summableI, of "{i}" f]] by simp lemma sums_ennreal[simp]: "(\i. 0 \ f i) \ 0 \ x \ (\i. ennreal (f i)) sums ennreal x \ f sums x" unfolding sums_def by (simp add: always_eventually sum_nonneg) lemma summable_suminf_not_top: "(\i. 0 \ f i) \ (\i. ennreal (f i)) \ top \ summable f" using summable_sums[OF summableI, of "\i. ennreal (f i)"] by (cases "\i. ennreal (f i)" rule: ennreal_cases) (auto simp: summable_def) lemma suminf_ennreal[simp]: "(\i. 0 \ f i) \ (\i. ennreal (f i)) \ top \ (\i. ennreal (f i)) = ennreal (\i. f i)" by (rule sums_unique[symmetric]) (simp add: summable_suminf_not_top suminf_nonneg summable_sums) lemma sums_enn2ereal[simp]: "(\i. enn2ereal (f i)) sums enn2ereal x \ f sums x" unfolding sums_def by (simp add: always_eventually sum_nonneg) lemma suminf_enn2ereal[simp]: "(\i. enn2ereal (f i)) = enn2ereal (suminf f)" by (rule sums_unique[symmetric]) (simp add: summable_sums) lemma transfer_e2ennreal_suminf [transfer_rule]: "rel_fun (rel_fun (=) pcr_ennreal) pcr_ennreal suminf suminf" by (auto simp: rel_funI rel_fun_eq_pcr_ennreal comp_def) lemma ennreal_suminf_cmult[simp]: "(\i. r * f i) = r * (\i. f i::ennreal)" by transfer (auto intro!: suminf_cmult_ereal) lemma ennreal_suminf_multc[simp]: "(\i. f i * r) = (\i. f i::ennreal) * r" using ennreal_suminf_cmult[of r f] by (simp add: ac_simps) lemma ennreal_suminf_divide[simp]: "(\i. f i / r) = (\i. f i::ennreal) / r" by (simp add: divide_ennreal_def) lemma ennreal_suminf_neq_top: "summable f \ (\i. 0 \ f i) \ (\i. ennreal (f i)) \ top" using sums_ennreal[of f "suminf f"] by (simp add: suminf_nonneg flip: sums_unique summable_sums_iff del: sums_ennreal) lemma suminf_ennreal_eq: "(\i. 0 \ f i) \ f sums x \ (\i. ennreal (f i)) = ennreal x" using suminf_nonneg[of f] sums_unique[of f x] by (intro sums_unique[symmetric]) (auto simp: summable_sums_iff) lemma ennreal_suminf_bound_add: fixes f :: "nat \ ennreal" shows "(\N. (\n x) \ suminf f + y \ x" by transfer (auto intro!: suminf_bound_add) lemma ennreal_suminf_SUP_eq_directed: fixes f :: "'a \ nat \ ennreal" assumes *: "\N i j. i \ I \ j \ I \ finite N \ \k\I. \n\N. f i n \ f k n \ f j n \ f k n" shows "(\n. SUP i\I. f i n) = (SUP i\I. \n. f i n)" proof cases assume "I \ {}" then obtain i where "i \ I" by auto from * show ?thesis by (transfer fixing: I) (auto simp: max_absorb2 SUP_upper2[OF \i \ I\] suminf_nonneg summable_ereal_pos \I \ {}\ intro!: suminf_SUP_eq_directed) qed (simp add: bot_ennreal) lemma INF_ennreal_add_const: fixes f g :: "nat \ ennreal" shows "(INF i. f i + c) = (INF i. f i) + c" using continuous_at_Inf_mono[of "\x. x + c" "f`UNIV"] using continuous_add[of "at_right (Inf (range f))", of "\x. x" "\x. c"] by (auto simp: mono_def image_comp) lemma INF_ennreal_const_add: fixes f g :: "nat \ ennreal" shows "(INF i. c + f i) = c + (INF i. f i)" using INF_ennreal_add_const[of f c] by (simp add: ac_simps) lemma SUP_mult_left_ennreal: "c * (SUP i\I. f i) = (SUP i\I. c * f i ::ennreal)" proof cases assume "I \ {}" then show ?thesis by transfer (auto simp add: SUP_ereal_mult_left max_absorb2 SUP_upper2) qed (simp add: bot_ennreal) lemma SUP_mult_right_ennreal: "(SUP i\I. f i) * c = (SUP i\I. f i * c ::ennreal)" using SUP_mult_left_ennreal by (simp add: mult.commute) lemma SUP_divide_ennreal: "(SUP i\I. f i) / c = (SUP i\I. f i / c ::ennreal)" using SUP_mult_right_ennreal by (simp add: divide_ennreal_def) lemma ennreal_SUP_of_nat_eq_top: "(SUP x. of_nat x :: ennreal) = top" proof (intro antisym top_greatest le_SUP_iff[THEN iffD2] allI impI) fix y :: ennreal assume "y < top" then obtain r where "y = ennreal r" by (cases y rule: ennreal_cases) auto then show "\i\UNIV. y < of_nat i" using reals_Archimedean2[of "max 1 r"] zero_less_one by (simp add: ennreal_Ex_less_of_nat) qed lemma ennreal_SUP_eq_top: fixes f :: "'a \ ennreal" assumes "\n. \i\I. of_nat n \ f i" shows "(SUP i \ I. f i) = top" proof - have "(SUP x. of_nat x :: ennreal) \ (SUP i \ I. f i)" using assms by (auto intro!: SUP_least intro: SUP_upper2) then show ?thesis by (auto simp: ennreal_SUP_of_nat_eq_top top_unique) qed lemma ennreal_INF_const_minus: fixes f :: "'a \ ennreal" shows "I \ {} \ (SUP x\I. c - f x) = c - (INF x\I. f x)" by (transfer fixing: I) (simp add: sup_max[symmetric] SUP_sup_const1 SUP_ereal_minus_right del: sup_ereal_def) lemma of_nat_Sup_ennreal: assumes "A \ {}" "bdd_above A" shows "of_nat (Sup A) = (SUP a\A. of_nat a :: ennreal)" proof (intro antisym) show "(SUP a\A. of_nat a::ennreal) \ of_nat (Sup A)" by (intro SUP_least of_nat_mono) (auto intro: cSup_upper assms) have "Sup A \ A" using assms by (auto simp: Sup_nat_def bdd_above_nat) then show "of_nat (Sup A) \ (SUP a\A. of_nat a::ennreal)" by (intro SUP_upper) qed lemma ennreal_tendsto_const_minus: fixes g :: "'a \ ennreal" assumes ae: "\\<^sub>F x in F. g x \ c" assumes g: "((\x. c - g x) \ 0) F" shows "(g \ c) F" proof (cases c rule: ennreal_cases) case top with tendsto_unique[OF _ g, of "top"] show ?thesis by (cases "F = bot") auto next case (real r) then have "\x. \q\0. g x \ c \ (g x = ennreal q \ q \ r)" by (auto simp: le_ennreal_iff) then obtain f where *: "0 \ f x" "g x = ennreal (f x)" "f x \ r" if "g x \ c" for x by metis from ae have ae2: "\\<^sub>F x in F. c - g x = ennreal (r - f x) \ f x \ r \ g x = ennreal (f x) \ 0 \ f x" proof eventually_elim fix x assume "g x \ c" with *[of x] \0 \ r\ show "c - g x = ennreal (r - f x) \ f x \ r \ g x = ennreal (f x) \ 0 \ f x" by (auto simp: real ennreal_minus) qed with g have "((\x. ennreal (r - f x)) \ ennreal 0) F" by (auto simp add: tendsto_cong eventually_conj_iff) with ae2 have "((\x. r - f x) \ 0) F" by (subst (asm) tendsto_ennreal_iff) (auto elim: eventually_mono) then have "(f \ r) F" by (rule Lim_transform2[OF tendsto_const]) with ae2 have "((\x. ennreal (f x)) \ ennreal r) F" by (subst tendsto_ennreal_iff) (auto elim: eventually_mono simp: real) with ae2 show ?thesis by (auto simp: real tendsto_cong eventually_conj_iff) qed lemma ennreal_SUP_add: fixes f g :: "nat \ ennreal" shows "incseq f \ incseq g \ (SUP i. f i + g i) = Sup (f ` UNIV) + Sup (g ` UNIV)" unfolding incseq_def le_fun_def by transfer (simp add: SUP_ereal_add incseq_def le_fun_def max_absorb2 SUP_upper2) lemma ennreal_SUP_sum: fixes f :: "'a \ nat \ ennreal" shows "(\i. i \ I \ incseq (f i)) \ (SUP n. \i\I. f i n) = (\i\I. SUP n. f i n)" unfolding incseq_def by transfer (simp add: SUP_ereal_sum incseq_def SUP_upper2 max_absorb2 sum_nonneg) lemma ennreal_liminf_minus: fixes f :: "nat \ ennreal" shows "(\n. f n \ c) \ liminf (\n. c - f n) = c - limsup f" apply transfer apply (simp add: ereal_diff_positive max.absorb2 liminf_ereal_cminus) apply (subst max.absorb2) apply (rule ereal_diff_positive) apply (rule Limsup_bounded) apply auto done lemma ennreal_continuous_on_cmult: "(c::ennreal) < top \ continuous_on A f \ continuous_on A (\x. c * f x)" by (transfer fixing: A) (auto intro: continuous_on_cmult_ereal) lemma ennreal_tendsto_cmult: "(c::ennreal) < top \ (f \ x) F \ ((\x. c * f x) \ c * x) F" by (rule continuous_on_tendsto_compose[where g=f, OF ennreal_continuous_on_cmult, where s=UNIV]) (auto simp: continuous_on_id) lemma tendsto_ennrealI[intro, simp, tendsto_intros]: "(f \ x) F \ ((\x. ennreal (f x)) \ ennreal x) F" by (auto simp: ennreal_def intro!: continuous_on_tendsto_compose[OF continuous_on_e2ennreal[of UNIV]] tendsto_max) lemma tendsto_enn2erealI [tendsto_intros]: assumes "(f \ l) F" shows "((\i. enn2ereal(f i)) \ enn2ereal l) F" using tendsto_enn2ereal_iff assms by auto lemma tendsto_e2ennrealI [tendsto_intros]: assumes "(f \ l) F" shows "((\i. e2ennreal(f i)) \ e2ennreal l) F" proof - have *: "e2ennreal (max x 0) = e2ennreal x" for x by (simp add: e2ennreal_def max.commute) have "((\i. max (f i) 0) \ max l 0) F" apply (intro tendsto_intros) using assms by auto then have "((\i. enn2ereal(e2ennreal (max (f i) 0))) \ enn2ereal (e2ennreal (max l 0))) F" by (subst enn2ereal_e2ennreal, auto)+ then have "((\i. e2ennreal (max (f i) 0)) \ e2ennreal (max l 0)) F" using tendsto_enn2ereal_iff by auto then show ?thesis unfolding * by auto qed lemma ennreal_suminf_minus: fixes f g :: "nat \ ennreal" shows "(\i. g i \ f i) \ suminf f \ top \ suminf g \ top \ (\i. f i - g i) = suminf f - suminf g" by transfer (auto simp add: max.absorb2 ereal_diff_positive suminf_le_pos top_ereal_def intro!: suminf_ereal_minus) lemma ennreal_Sup_countable_SUP: "A \ {} \ \f::nat \ ennreal. incseq f \ range f \ A \ Sup A = (SUP i. f i)" unfolding incseq_def apply transfer subgoal for A using Sup_countable_SUP[of A] by (force simp add: incseq_def[symmetric] SUP_upper2 max.absorb2 image_subset_iff Sup_upper2 cong: conj_cong) done lemma ennreal_Inf_countable_INF: "A \ {} \ \f::nat \ ennreal. decseq f \ range f \ A \ Inf A = (INF i. f i)" including ennreal.lifting unfolding decseq_def apply transfer subgoal for A using Inf_countable_INF[of A] apply (clarsimp simp flip: decseq_def) subgoal for f by (intro exI[of _ f]) auto done done lemma ennreal_SUP_countable_SUP: "A \ {} \ \f::nat \ ennreal. range f \ g`A \ Sup (g ` A) = Sup (f ` UNIV)" using ennreal_Sup_countable_SUP [of "g`A"] by auto lemma of_nat_tendsto_top_ennreal: "(\n::nat. of_nat n :: ennreal) \ top" using LIMSEQ_SUP[of "of_nat :: nat \ ennreal"] by (simp add: ennreal_SUP_of_nat_eq_top incseq_def) lemma SUP_sup_continuous_ennreal: fixes f :: "ennreal \ 'a::complete_lattice" assumes f: "sup_continuous f" and "I \ {}" shows "(SUP i\I. f (g i)) = f (SUP i\I. g i)" proof (rule antisym) show "(SUP i\I. f (g i)) \ f (SUP i\I. g i)" by (rule mono_SUP[OF sup_continuous_mono[OF f]]) from ennreal_Sup_countable_SUP[of "g`I"] \I \ {}\ obtain M :: "nat \ ennreal" where "incseq M" and M: "range M \ g ` I" and eq: "(SUP i \ I. g i) = (SUP i. M i)" by auto have "f (SUP i \ I. g i) = (SUP i \ range M. f i)" unfolding eq sup_continuousD[OF f \mono M\] by (simp add: image_comp) also have "\ \ (SUP i \ I. f (g i))" by (insert M, drule SUP_subset_mono) (auto simp add: image_comp) finally show "f (SUP i \ I. g i) \ (SUP i \ I. f (g i))" . qed lemma ennreal_suminf_SUP_eq: fixes f :: "nat \ nat \ ennreal" shows "(\i. incseq (\n. f n i)) \ (\i. SUP n. f n i) = (SUP n. \i. f n i)" apply (rule ennreal_suminf_SUP_eq_directed) subgoal for N n j by (auto simp: incseq_def intro!:exI[of _ "max n j"]) done lemma ennreal_SUP_add_left: fixes c :: ennreal shows "I \ {} \ (SUP i\I. f i + c) = (SUP i\I. f i) + c" apply transfer apply (simp add: SUP_ereal_add_left) apply (subst (1 2) max.absorb2) apply (auto intro: SUP_upper2 add_nonneg_nonneg) done lemma ennreal_SUP_const_minus: fixes f :: "'a \ ennreal" shows "I \ {} \ c < top \ (INF x\I. c - f x) = c - (SUP x\I. f x)" apply (transfer fixing: I) unfolding ex_in_conv[symmetric] apply (auto simp add: SUP_upper2 sup_absorb2 simp flip: sup_ereal_def) apply (subst INF_ereal_minus_right[symmetric]) apply (auto simp del: sup_ereal_def simp add: sup_INF) done subsection \Approximation lemmas\ lemma INF_approx_ennreal: fixes x::ennreal and e::real assumes "e > 0" assumes INF: "x = (INF i \ A. f i)" assumes "x \ \" shows "\i \ A. f i < x + e" proof - have "(INF i \ A. f i) < x + e" unfolding INF[symmetric] using \0 \x \ \\ by (cases x) auto then show ?thesis unfolding INF_less_iff . qed lemma SUP_approx_ennreal: fixes x::ennreal and e::real assumes "e > 0" "A \ {}" assumes SUP: "x = (SUP i \ A. f i)" assumes "x \ \" shows "\i \ A. x < f i + e" proof - have "x < x + e" using \0 \x \ \\ by (cases x) auto also have "x + e = (SUP i \ A. f i + e)" unfolding SUP ennreal_SUP_add_left[OF \A \ {}\] .. finally show ?thesis unfolding less_SUP_iff . qed lemma ennreal_approx_SUP: fixes x::ennreal assumes f_bound: "\i. i \ A \ f i \ x" assumes approx: "\e. (e::real) > 0 \ \i \ A. x \ f i + e" shows "x = (SUP i \ A. f i)" proof (rule antisym) show "x \ (SUP i\A. f i)" proof (rule ennreal_le_epsilon) fix e :: real assume "0 < e" - from approx[OF this] guess i .. - then have "x \ f i + e" + from approx[OF this] obtain i where "i \ A" and *: "x \ f i + ennreal e" + by blast + from * have "x \ f i + e" by simp also have "\ \ (SUP i\A. f i) + e" by (intro add_mono \i \ A\ SUP_upper order_refl) finally show "x \ (SUP i\A. f i) + e" . qed qed (intro SUP_least f_bound) lemma ennreal_approx_INF: fixes x::ennreal assumes f_bound: "\i. i \ A \ x \ f i" assumes approx: "\e. (e::real) > 0 \ \i \ A. f i \ x + e" shows "x = (INF i \ A. f i)" proof (rule antisym) show "(INF i\A. f i) \ x" proof (rule ennreal_le_epsilon) fix e :: real assume "0 < e" - from approx[OF this] guess i .. note i = this + from approx[OF this] obtain i where "i\A" "f i \ x + ennreal e" + by blast then have "(INF i\A. f i) \ f i" by (intro INF_lower) also have "\ \ x + e" by fact finally show "(INF i\A. f i) \ x + e" . qed qed (intro INF_greatest f_bound) lemma ennreal_approx_unit: "(\a::ennreal. 0 < a \ a < 1 \ a * z \ y) \ z \ y" apply (subst SUP_mult_right_ennreal[of "\x. x" "{0 <..< 1}" z, simplified]) apply (auto intro: SUP_least) done lemma suminf_ennreal2: "(\i. 0 \ f i) \ summable f \ (\i. ennreal (f i)) = ennreal (\i. f i)" using suminf_ennreal_eq by blast lemma less_top_ennreal: "x < top \ (\r\0. x = ennreal r)" by (cases x) auto lemma enn2real_less_iff[simp]: "x < top \ enn2real x < c \ x < c" using ennreal_less_iff less_top_ennreal by auto lemma enn2real_le_iff[simp]: "\x < top; c > 0\ \ enn2real x \ c \ x \ c" by (cases x) auto lemma enn2real_less: assumes "enn2real e < r" "e \ top" shows "e < ennreal r" using enn2real_less_iff assms top.not_eq_extremum by blast lemma enn2real_le: assumes "enn2real e \ r" "e \ top" shows "e \ ennreal r" by (metis assms enn2real_less ennreal_enn2real_if eq_iff less_le) lemma tendsto_top_iff_ennreal: fixes f :: "'a \ ennreal" shows "(f \ top) F \ (\l\0. eventually (\x. ennreal l < f x) F)" by (auto simp: less_top_ennreal order_tendsto_iff ) lemma ennreal_tendsto_top_eq_at_top: "((\z. ennreal (f z)) \ top) F \ (LIM z F. f z :> at_top)" unfolding filterlim_at_top_dense tendsto_top_iff_ennreal apply (auto simp: ennreal_less_iff) subgoal for y by (auto elim!: eventually_mono allE[of _ "max 0 y"]) done lemma tendsto_0_if_Limsup_eq_0_ennreal: fixes f :: "_ \ ennreal" shows "Limsup F f = 0 \ (f \ 0) F" using Liminf_le_Limsup[of F f] tendsto_iff_Liminf_eq_Limsup[of F f 0] by (cases "F = bot") auto lemma diff_le_self_ennreal[simp]: "a - b \ (a::ennreal)" by (cases a b rule: ennreal2_cases) (auto simp: ennreal_minus) lemma ennreal_ineq_diff_add: "b \ a \ a = b + (a - b::ennreal)" by transfer (auto simp: ereal_diff_positive max.absorb2 ereal_ineq_diff_add) lemma ennreal_mult_strict_left_mono: "(a::ennreal) < c \ 0 < b \ b < top \ b * a < b * c" by transfer (auto intro!: ereal_mult_strict_left_mono) lemma ennreal_between: "0 < e \ 0 < x \ x < top \ x - e < (x::ennreal)" by transfer (auto intro!: ereal_between) lemma minus_less_iff_ennreal: "b < top \ b \ a \ a - b < c \ a < c + (b::ennreal)" by transfer (auto simp: top_ereal_def ereal_minus_less le_less) lemma tendsto_zero_ennreal: assumes ev: "\r. 0 < r \ \\<^sub>F x in F. f x < ennreal r" shows "(f \ 0) F" proof (rule order_tendstoI) fix e::ennreal assume "e > 0" obtain e'::real where "e' > 0" "ennreal e' < e" using \0 < e\ dense[of 0 "if e = top then 1 else (enn2real e)"] by (cases e) (auto simp: ennreal_less_iff) from ev[OF \e' > 0\] show "\\<^sub>F x in F. f x < e" by eventually_elim (insert \ennreal e' < e\, auto) qed simp lifting_update ennreal.lifting lifting_forget ennreal.lifting subsection \\<^typ>\ennreal\ theorems\ lemma neq_top_trans: fixes x y :: ennreal shows "\ y \ top; x \ y \ \ x \ top" by (auto simp: top_unique) lemma diff_diff_ennreal: fixes a b :: ennreal shows "a \ b \ b \ \ \ b - (b - a) = a" by (cases a b rule: ennreal2_cases) (auto simp: ennreal_minus top_unique) lemma ennreal_less_one_iff[simp]: "ennreal x < 1 \ x < 1" by (cases "0 \ x") (auto simp: ennreal_neg ennreal_less_iff simp flip: ennreal_1) lemma SUP_const_minus_ennreal: fixes f :: "'a \ ennreal" shows "I \ {} \ (SUP x\I. c - f x) = c - (INF x\I. f x)" including ennreal.lifting by (transfer fixing: I) (simp add: SUP_sup_distrib[symmetric] SUP_ereal_minus_right flip: sup_ereal_def) lemma zero_minus_ennreal[simp]: "0 - (a::ennreal) = 0" including ennreal.lifting by transfer (simp split: split_max) lemma diff_diff_commute_ennreal: fixes a b c :: ennreal shows "a - b - c = a - c - b" by (cases a b c rule: ennreal3_cases) (simp_all add: ennreal_minus field_simps) lemma diff_gr0_ennreal: "b < (a::ennreal) \ 0 < a - b" including ennreal.lifting by transfer (auto simp: ereal_diff_gr0 ereal_diff_positive split: split_max) lemma divide_le_posI_ennreal: fixes x y z :: ennreal shows "x > 0 \ z \ x * y \ z / x \ y" by (cases x y z rule: ennreal3_cases) (auto simp: divide_ennreal ennreal_mult[symmetric] field_simps top_unique) lemma add_diff_eq_ennreal: fixes x y z :: ennreal shows "z \ y \ x + (y - z) = x + y - z" including ennreal.lifting by transfer (insert add_mono[of "0::ereal"], auto simp add: ereal_diff_positive max.absorb2 add_diff_eq_ereal) lemma add_diff_inverse_ennreal: fixes x y :: ennreal shows "x \ y \ x + (y - x) = y" by (cases x) (simp_all add: top_unique add_diff_eq_ennreal) lemma add_diff_eq_iff_ennreal[simp]: fixes x y :: ennreal shows "x + (y - x) = y \ x \ y" proof assume *: "x + (y - x) = y" show "x \ y" by (subst *[symmetric]) simp qed (simp add: add_diff_inverse_ennreal) lemma add_diff_le_ennreal: "a + b - c \ a + (b - c::ennreal)" apply (cases a b c rule: ennreal3_cases) subgoal for a' b' c' by (cases "0 \ b' - c'") (simp_all add: ennreal_minus top_add ennreal_neg flip: ennreal_plus) apply (simp_all add: top_add flip: ennreal_plus) done lemma diff_eq_0_ennreal: "a < top \ a \ b \ a - b = (0::ennreal)" using ennreal_minus_pos_iff gr_zeroI not_less by blast lemma diff_diff_ennreal': fixes x y z :: ennreal shows "z \ y \ y - z \ x \ x - (y - z) = x + z - y" by (cases x; cases y; cases z) (auto simp add: top_add add_top minus_top_ennreal ennreal_minus top_unique simp flip: ennreal_plus) lemma diff_diff_ennreal'': fixes x y z :: ennreal shows "z \ y \ x - (y - z) = (if y - z \ x then x + z - y else 0)" by (cases x; cases y; cases z) (auto simp add: top_add add_top minus_top_ennreal ennreal_minus top_unique ennreal_neg simp flip: ennreal_plus) lemma power_less_top_ennreal: fixes x :: ennreal shows "x ^ n < top \ x < top \ n = 0" using power_eq_top_ennreal[of x n] by (auto simp: less_top) lemma ennreal_divide_times: "(a / b) * c = a * (c / b :: ennreal)" by (simp add: mult.commute ennreal_times_divide) lemma diff_less_top_ennreal: "a - b < top \ a < (top :: ennreal)" by (cases a; cases b) (auto simp: ennreal_minus) lemma divide_less_ennreal: "b \ 0 \ b < top \ a / b < c \ a < (c * b :: ennreal)" by (cases a; cases b; cases c) (auto simp: divide_ennreal ennreal_mult[symmetric] ennreal_less_iff field_simps ennreal_top_mult ennreal_top_divide) lemma one_less_numeral[simp]: "1 < (numeral n::ennreal) \ (num.One < n)" by (simp flip: ennreal_1 ennreal_numeral add: ennreal_less_iff) lemma divide_eq_1_ennreal: "a / b = (1::ennreal) \ (b \ top \ b \ 0 \ b = a)" by (cases a ; cases b; cases "b = 0") (auto simp: ennreal_top_divide divide_ennreal split: if_split_asm) lemma ennreal_mult_cancel_left: "(a * b = a * c) = (a = top \ b \ 0 \ c \ 0 \ a = 0 \ b = (c::ennreal))" by (cases a; cases b; cases c) (auto simp: ennreal_mult[symmetric] ennreal_mult_top ennreal_top_mult) lemma ennreal_minus_if: "ennreal a - ennreal b = ennreal (if 0 \ b then (if b \ a then a - b else 0) else a)" by (auto simp: ennreal_minus ennreal_neg) lemma ennreal_plus_if: "ennreal a + ennreal b = ennreal (if 0 \ a then (if 0 \ b then a + b else a) else b)" by (auto simp: ennreal_neg) lemma power_le_one_iff: "0 \ (a::real) \ a ^ n \ 1 \ (n = 0 \ a \ 1)" by (metis (mono_tags, opaque_lifting) le_less neq0_conv not_le one_le_power power_0 power_eq_imp_eq_base power_le_one zero_le_one) lemma ennreal_diff_le_mono_left: "a \ b \ a - c \ (b::ennreal)" using ennreal_mono_minus[of 0 c a, THEN order_trans, of b] by simp lemma ennreal_minus_le_iff: "a - b \ c \ (a \ b + (c::ennreal) \ (a = top \ b = top \ c = top))" by (cases a; cases b; cases c) (auto simp: top_unique top_add add_top ennreal_minus simp flip: ennreal_plus) lemma ennreal_le_minus_iff: "a \ b - c \ (a + c \ (b::ennreal) \ (a = 0 \ b \ c))" by (cases a; cases b; cases c) (auto simp: top_unique top_add add_top ennreal_minus ennreal_le_iff2 simp flip: ennreal_plus) lemma diff_add_eq_diff_diff_swap_ennreal: "x - (y + z :: ennreal) = x - y - z" by (cases x; cases y; cases z) (auto simp: ennreal_minus_if add_top top_add simp flip: ennreal_plus) lemma diff_add_assoc2_ennreal: "b \ a \ (a - b + c::ennreal) = a + c - b" by (cases a; cases b; cases c) (auto simp add: ennreal_minus_if ennreal_plus_if add_top top_add top_unique simp del: ennreal_plus) lemma diff_gt_0_iff_gt_ennreal: "0 < a - b \ (a = top \ b = top \ b < (a::ennreal))" by (cases a; cases b) (auto simp: ennreal_minus_if ennreal_less_iff) lemma diff_eq_0_iff_ennreal: "(a - b::ennreal) = 0 \ (a < top \ a \ b)" by (cases a) (auto simp: ennreal_minus_eq_0 diff_eq_0_ennreal) lemma add_diff_self_ennreal: "a + (b - a::ennreal) = (if a \ b then b else a)" by (auto simp: diff_eq_0_iff_ennreal less_top) lemma diff_add_self_ennreal: "(b - a + a::ennreal) = (if a \ b then b else a)" by (auto simp: diff_add_cancel_ennreal diff_eq_0_iff_ennreal less_top) lemma ennreal_minus_cancel_iff: fixes a b c :: ennreal shows "a - b = a - c \ (b = c \ (a \ b \ a \ c) \ a = top)" by (cases a; cases b; cases c) (auto simp: ennreal_minus_if) text \The next lemma is wrong for $a = top$, for $b = c = 1$ for instance.\ lemma ennreal_right_diff_distrib: fixes a b c :: ennreal assumes "a \ top" shows "a * (b - c) = a * b - a * c" apply (cases a; cases b; cases c) apply (use assms in \auto simp add: ennreal_mult_top ennreal_minus ennreal_mult' [symmetric]\) apply (simp add: algebra_simps) done lemma SUP_diff_ennreal: "c < top \ (SUP i\I. f i - c :: ennreal) = (SUP i\I. f i) - c" by (auto intro!: SUP_eqI ennreal_minus_mono SUP_least intro: SUP_upper simp: ennreal_minus_cancel_iff ennreal_minus_le_iff less_top[symmetric]) lemma ennreal_SUP_add_right: fixes c :: ennreal shows "I \ {} \ c + (SUP i\I. f i) = (SUP i\I. c + f i)" using ennreal_SUP_add_left[of I f c] by (simp add: add.commute) lemma SUP_add_directed_ennreal: fixes f g :: "_ \ ennreal" assumes directed: "\i j. i \ I \ j \ I \ \k\I. f i + g j \ f k + g k" shows "(SUP i\I. f i + g i) = (SUP i\I. f i) + (SUP i\I. g i)" proof cases assume "I = {}" then show ?thesis by (simp add: bot_ereal_def) next assume "I \ {}" show ?thesis proof (rule antisym) show "(SUP i\I. f i + g i) \ (SUP i\I. f i) + (SUP i\I. g i)" by (rule SUP_least; intro add_mono SUP_upper) next have "(SUP i\I. f i) + (SUP i\I. g i) = (SUP i\I. f i + (SUP i\I. g i))" by (intro ennreal_SUP_add_left[symmetric] \I \ {}\) also have "\ = (SUP i\I. (SUP j\I. f i + g j))" using \I \ {}\ by (simp add: ennreal_SUP_add_right) also have "\ \ (SUP i\I. f i + g i)" using directed by (intro SUP_least) (blast intro: SUP_upper2) finally show "(SUP i\I. f i) + (SUP i\I. g i) \ (SUP i\I. f i + g i)" . qed qed lemma enn2real_eq_0_iff: "enn2real x = 0 \ x = 0 \ x = top" by (cases x) auto lemma continuous_on_diff_ennreal: "continuous_on A f \ continuous_on A g \ (\x. x \ A \ f x \ top) \ (\x. x \ A \ g x \ top) \ continuous_on A (\z. f z - g z::ennreal)" including ennreal.lifting proof (transfer fixing: A, simp add: top_ereal_def) fix f g :: "'a \ ereal" assume "\x. 0 \ f x" "\x. 0 \ g x" "continuous_on A f" "continuous_on A g" moreover assume "f x \ \" "g x \ \" if "x \ A" for x ultimately show "continuous_on A (\z. max 0 (f z - g z))" by (intro continuous_on_max continuous_on_const continuous_on_diff_ereal) auto qed lemma tendsto_diff_ennreal: "(f \ x) F \ (g \ y) F \ x \ top \ y \ top \ ((\z. f z - g z::ennreal) \ x - y) F" using continuous_on_tendsto_compose[where f="\x. fst x - snd x::ennreal" and s="{(x, y). x \ top \ y \ top}" and g="\x. (f x, g x)" and l="(x, y)" and F="F", OF continuous_on_diff_ennreal] by (auto simp: tendsto_Pair eventually_conj_iff less_top order_tendstoD continuous_on_fst continuous_on_snd continuous_on_id) declare lim_real_of_ereal [tendsto_intros] lemma tendsto_enn2real [tendsto_intros]: assumes "(u \ ennreal l) F" "l \ 0" shows "((\n. enn2real (u n)) \ l) F" unfolding enn2real_def by (metis assms enn2ereal_ennreal lim_real_of_ereal tendsto_enn2erealI) end diff --git a/src/HOL/Library/Extended_Real.thy b/src/HOL/Library/Extended_Real.thy --- a/src/HOL/Library/Extended_Real.thy +++ b/src/HOL/Library/Extended_Real.thy @@ -1,4200 +1,4200 @@ (* Title: HOL/Library/Extended_Real.thy Author: Johannes Hölzl, TU München Author: Robert Himmelmann, TU München Author: Armin Heller, TU München Author: Bogdan Grechuk, University of Edinburgh Author: Manuel Eberl, TU München *) 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 {.. 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 \\<^sub>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 \<^file>\~~/src/HOL/Analysis/Extended_Real_Limits.thy\. \ subsection \Definition and basic properties\ datatype ereal = ereal real | PInfty | MInfty lemma ereal_cong: "x = y \ ereal x = ereal y" by simp instantiation ereal :: uminus begin fun uminus_ereal where "- (ereal r) = ereal (- r)" | "- PInfty = MInfty" | "- MInfty = PInfty" instance .. end instantiation ereal :: infinity begin definition "(\::ereal) = PInfty" instance .. end declare [[coercion "ereal :: real \ ereal"]] lemma ereal_uminus_uminus[simp]: fixes a :: ereal shows "- (- a) = a" by (cases a) simp_all lemma shows PInfty_eq_infinity[simp]: "PInfty = \" and MInfty_eq_minfinity[simp]: "MInfty = - \" and MInfty_neq_PInfty[simp]: "\ \ - (\::ereal)" "- \ \ (\::ereal)" and MInfty_neq_ereal[simp]: "ereal r \ - \" "- \ \ ereal r" and PInfty_neq_ereal[simp]: "ereal r \ \" "\ \ ereal r" and PInfty_cases[simp]: "(case \ of ereal r \ f r | PInfty \ y | MInfty \ z) = y" and MInfty_cases[simp]: "(case - \ of ereal r \ f r | PInfty \ y | MInfty \ z) = z" by (simp_all add: infinity_ereal_def) declare PInfty_eq_infinity[code_post] MInfty_eq_minfinity[code_post] lemma [code_unfold]: "\ = PInfty" "- PInfty = MInfty" by simp_all lemma inj_ereal[simp]: "inj_on ereal A" unfolding inj_on_def by auto lemma ereal_cases[cases type: ereal]: obtains (real) r where "x = ereal r" | (PInf) "x = \" | (MInf) "x = -\" by (cases x) auto lemmas ereal2_cases = ereal_cases[case_product ereal_cases] lemmas ereal3_cases = ereal2_cases[case_product ereal_cases] lemma ereal_all_split: "\P. (\x::ereal. P x) \ P \ \ (\x. P (ereal x)) \ P (-\)" by (metis ereal_cases) lemma ereal_ex_split: "\P. (\x::ereal. P x) \ P \ \ (\x. P (ereal x)) \ P (-\)" by (metis ereal_cases) lemma ereal_uminus_eq_iff[simp]: fixes a b :: ereal shows "-a = -b \ a = b" by (cases rule: ereal2_cases[of a b]) simp_all function real_of_ereal :: "ereal \ real" where "real_of_ereal (ereal r) = r" | "real_of_ereal \ = 0" | "real_of_ereal (-\) = 0" by (auto intro: ereal_cases) termination by standard (rule wf_empty) lemma real_of_ereal[simp]: "real_of_ereal (- x :: ereal) = - (real_of_ereal x)" by (cases x) simp_all lemma range_ereal[simp]: "range ereal = UNIV - {\, -\}" proof safe fix x assume "x \ range ereal" "x \ \" then show "x = -\" by (cases x) auto qed auto lemma ereal_range_uminus[simp]: "range uminus = (UNIV::ereal set)" proof safe fix x :: ereal show "x \ range uminus" by (intro image_eqI[of _ _ "-x"]) auto qed auto instantiation ereal :: abs begin function abs_ereal where "\ereal r\ = ereal \r\" | "\-\\ = (\::ereal)" | "\\\ = (\::ereal)" by (auto intro: ereal_cases) termination proof qed (rule wf_empty) instance .. end lemma abs_eq_infinity_cases[elim!]: fixes x :: ereal assumes "\x\ = \" obtains "x = \" | "x = -\" using assms by (cases x) auto lemma abs_neq_infinity_cases[elim!]: fixes x :: ereal assumes "\x\ \ \" obtains r where "x = ereal r" using assms by (cases x) auto lemma abs_ereal_uminus[simp]: fixes x :: ereal shows "\- x\ = \x\" by (cases x) auto lemma ereal_infinity_cases: fixes a :: ereal shows "a \ \ \ a \ -\ \ \a\ \ \" by auto subsubsection "Addition" instantiation ereal :: "{one,comm_monoid_add,zero_neq_one}" begin definition "0 = ereal 0" definition "1 = ereal 1" function plus_ereal where "ereal r + ereal p = ereal (r + p)" | "\ + a = (\::ereal)" | "a + \ = (\::ereal)" | "ereal r + -\ = - \" | "-\ + ereal p = -(\::ereal)" | "-\ + -\ = -(\::ereal)" proof goal_cases case prems: (1 P x) then obtain a b where "x = (a, b)" by (cases x) auto with prems show P by (cases rule: ereal2_cases[of a b]) auto qed auto termination by standard (rule wf_empty) lemma Infty_neq_0[simp]: "(\::ereal) \ 0" "0 \ (\::ereal)" "-(\::ereal) \ 0" "0 \ -(\::ereal)" by (simp_all add: zero_ereal_def) lemma ereal_eq_0[simp]: "ereal r = 0 \ r = 0" "0 = ereal r \ r = 0" unfolding zero_ereal_def by simp_all lemma ereal_eq_1[simp]: "ereal r = 1 \ r = 1" "1 = ereal r \ r = 1" unfolding one_ereal_def by simp_all instance proof fix a b c :: ereal show "0 + a = a" by (cases a) (simp_all add: zero_ereal_def) show "a + b = b + a" by (cases rule: ereal2_cases[of a b]) simp_all show "a + b + c = a + (b + c)" by (cases rule: ereal3_cases[of a b c]) simp_all show "0 \ (1::ereal)" by (simp add: one_ereal_def zero_ereal_def) qed end lemma ereal_0_plus [simp]: "ereal 0 + x = x" and plus_ereal_0 [simp]: "x + ereal 0 = x" by(simp_all flip: zero_ereal_def) instance ereal :: numeral .. lemma real_of_ereal_0[simp]: "real_of_ereal (0::ereal) = 0" unfolding zero_ereal_def by simp lemma abs_ereal_zero[simp]: "\0\ = (0::ereal)" unfolding zero_ereal_def abs_ereal.simps by simp lemma ereal_uminus_zero[simp]: "- 0 = (0::ereal)" by (simp add: zero_ereal_def) lemma ereal_uminus_zero_iff[simp]: fixes a :: ereal shows "-a = 0 \ a = 0" by (cases a) simp_all lemma ereal_plus_eq_PInfty[simp]: fixes a b :: ereal shows "a + b = \ \ a = \ \ b = \" by (cases rule: ereal2_cases[of a b]) auto lemma ereal_plus_eq_MInfty[simp]: fixes a b :: ereal shows "a + b = -\ \ (a = -\ \ b = -\) \ a \ \ \ b \ \" by (cases rule: ereal2_cases[of a b]) auto lemma ereal_add_cancel_left: fixes a b :: ereal assumes "a \ -\" shows "a + b = a + c \ a = \ \ b = c" using assms by (cases rule: ereal3_cases[of a b c]) auto lemma ereal_add_cancel_right: fixes a b :: ereal assumes "a \ -\" shows "b + a = c + a \ a = \ \ b = c" using assms by (cases rule: ereal3_cases[of a b c]) auto lemma ereal_real: "ereal (real_of_ereal x) = (if \x\ = \ then 0 else x)" by (cases x) simp_all lemma real_of_ereal_add: fixes a b :: ereal shows "real_of_ereal (a + b) = (if (\a\ = \) \ (\b\ = \) \ (\a\ \ \) \ (\b\ \ \) then real_of_ereal a + real_of_ereal b else 0)" by (cases rule: ereal2_cases[of a b]) auto subsubsection "Linear order on \<^typ>\ereal\" instantiation ereal :: linorder begin function less_ereal where " ereal x < ereal y \ x < y" | "(\::ereal) < a \ False" | " a < -(\::ereal) \ False" | "ereal x < \ \ True" | " -\ < ereal r \ True" | " -\ < (\::ereal) \ True" proof goal_cases case prems: (1 P x) then obtain a b where "x = (a,b)" by (cases x) auto with prems show P by (cases rule: ereal2_cases[of a b]) auto qed simp_all termination by (relation "{}") simp definition "x \ (y::ereal) \ x < y \ x = y" lemma ereal_infty_less[simp]: fixes x :: ereal shows "x < \ \ (x \ \)" "-\ < x \ (x \ -\)" by (cases x, simp_all) (cases x, simp_all) lemma ereal_infty_less_eq[simp]: fixes x :: ereal shows "\ \ x \ x = \" and "x \ -\ \ x = -\" by (auto simp add: less_eq_ereal_def) lemma ereal_less[simp]: "ereal r < 0 \ (r < 0)" "0 < ereal r \ (0 < r)" "ereal r < 1 \ (r < 1)" "1 < ereal r \ (1 < r)" "0 < (\::ereal)" "-(\::ereal) < 0" by (simp_all add: zero_ereal_def one_ereal_def) lemma ereal_less_eq[simp]: "x \ (\::ereal)" "-(\::ereal) \ x" "ereal r \ ereal p \ r \ p" "ereal r \ 0 \ r \ 0" "0 \ ereal r \ 0 \ r" "ereal r \ 1 \ r \ 1" "1 \ ereal r \ 1 \ r" by (auto simp add: less_eq_ereal_def zero_ereal_def one_ereal_def) lemma ereal_infty_less_eq2: "a \ b \ a = \ \ b = (\::ereal)" "a \ b \ b = -\ \ a = -(\::ereal)" by simp_all instance proof fix x y z :: ereal show "x \ x" by (cases x) simp_all show "x < y \ x \ y \ \ y \ x" by (cases rule: ereal2_cases[of x y]) auto show "x \ y \ y \ x " by (cases rule: ereal2_cases[of x y]) auto { assume "x \ y" "y \ x" then show "x = y" by (cases rule: ereal2_cases[of x y]) auto } { assume "x \ y" "y \ z" then show "x \ z" by (cases rule: ereal3_cases[of x y z]) auto } qed end lemma ereal_dense2: "x < y \ \z. x < ereal z \ ereal z < y" using lt_ex gt_ex dense by (cases x y rule: ereal2_cases) auto instance ereal :: dense_linorder by standard (blast dest: ereal_dense2) instance ereal :: ordered_comm_monoid_add proof fix a b c :: ereal assume "a \ b" then show "c + a \ c + b" by (cases rule: ereal3_cases[of a b c]) auto qed lemma ereal_one_not_less_zero_ereal[simp]: "\ 1 < (0::ereal)" by (simp add: zero_ereal_def) lemma real_of_ereal_positive_mono: fixes x y :: ereal shows "0 \ x \ x \ y \ y \ \ \ real_of_ereal x \ real_of_ereal y" by (cases rule: ereal2_cases[of x y]) auto lemma ereal_MInfty_lessI[intro, simp]: fixes a :: ereal shows "a \ -\ \ -\ < a" by (cases a) auto lemma ereal_less_PInfty[intro, simp]: fixes a :: ereal shows "a \ \ \ a < \" by (cases a) auto lemma ereal_less_ereal_Ex: fixes a b :: ereal shows "x < ereal r \ x = -\ \ (\p. p < r \ x = ereal p)" by (cases x) auto lemma less_PInf_Ex_of_nat: "x \ \ \ (\n::nat. x < ereal (real n))" proof (cases x) case (real r) then show ?thesis using reals_Archimedean2[of r] by simp qed simp_all lemma ereal_add_strict_mono2: fixes a b c d :: ereal assumes "a < b" and "c < d" shows "a + c < b + d" using assms by (cases a; force simp add: elim: less_ereal.elims) lemma ereal_minus_le_minus[simp]: fixes a b :: ereal shows "- a \ - b \ b \ a" by (cases rule: ereal2_cases[of a b]) auto lemma ereal_minus_less_minus[simp]: fixes a b :: ereal shows "- a < - b \ b < a" by (cases rule: ereal2_cases[of a b]) auto lemma ereal_le_real_iff: "x \ real_of_ereal y \ (\y\ \ \ \ ereal x \ y) \ (\y\ = \ \ x \ 0)" by (cases y) auto lemma real_le_ereal_iff: "real_of_ereal y \ x \ (\y\ \ \ \ y \ ereal x) \ (\y\ = \ \ 0 \ x)" by (cases y) auto lemma ereal_less_real_iff: "x < real_of_ereal y \ (\y\ \ \ \ ereal x < y) \ (\y\ = \ \ x < 0)" by (cases y) auto lemma real_less_ereal_iff: "real_of_ereal y < x \ (\y\ \ \ \ y < ereal x) \ (\y\ = \ \ 0 < x)" by (cases y) auto text \ To help with inferences like \<^prop>\a < ereal x \ x < y \ a < ereal y\, where x and y are real. \ lemma le_ereal_le: "a \ ereal x \ x \ y \ a \ ereal y" using ereal_less_eq(3) order.trans by blast lemma le_ereal_less: "a \ ereal x \ x < y \ a < ereal y" by (simp add: le_less_trans) lemma less_ereal_le: "a < ereal x \ x \ y \ a < ereal y" using ereal_less_ereal_Ex by auto lemma ereal_le_le: "ereal y \ a \ x \ y \ ereal x \ a" by (simp add: order_subst2) lemma ereal_le_less: "ereal y \ a \ x < y \ ereal x < a" by (simp add: dual_order.strict_trans1) lemma ereal_less_le: "ereal y < a \ x \ y \ ereal x < a" using ereal_less_eq(3) le_less_trans by blast lemma real_of_ereal_pos: fixes x :: ereal shows "0 \ x \ 0 \ real_of_ereal x" by (cases x) auto lemmas real_of_ereal_ord_simps = ereal_le_real_iff real_le_ereal_iff ereal_less_real_iff real_less_ereal_iff lemma abs_ereal_ge0[simp]: "0 \ x \ \x :: ereal\ = x" by (cases x) auto lemma abs_ereal_less0[simp]: "x < 0 \ \x :: ereal\ = -x" by (cases x) auto lemma abs_ereal_pos[simp]: "0 \ \x :: ereal\" by (cases x) auto lemma ereal_abs_leI: fixes x y :: ereal shows "\ x \ y; -x \ y \ \ \x\ \ y" by(cases x y rule: ereal2_cases)(simp_all) lemma ereal_abs_add: fixes a b::ereal shows "abs(a+b) \ abs a + abs b" by (cases rule: ereal2_cases[of a b]) (auto) lemma real_of_ereal_le_0[simp]: "real_of_ereal (x :: ereal) \ 0 \ x \ 0 \ x = \" by (cases x) auto lemma abs_real_of_ereal[simp]: "\real_of_ereal (x :: ereal)\ = real_of_ereal \x\" by (cases x) auto lemma zero_less_real_of_ereal: fixes x :: ereal shows "0 < real_of_ereal x \ 0 < x \ x \ \" by (cases x) auto lemma ereal_0_le_uminus_iff[simp]: fixes a :: ereal shows "0 \ - a \ a \ 0" by (cases rule: ereal2_cases[of a]) auto lemma ereal_uminus_le_0_iff[simp]: fixes a :: ereal shows "- a \ 0 \ 0 \ a" by (cases rule: ereal2_cases[of a]) auto lemma ereal_add_strict_mono: fixes a b c d :: ereal assumes "a \ b" and "0 \ a" and "a \ \" and "c < d" shows "a + c < b + d" using assms by (cases rule: ereal3_cases[case_product ereal_cases, of a b c d]) auto lemma ereal_less_add: fixes a b c :: ereal shows "\a\ \ \ \ c < b \ a + c < a + b" by (cases rule: ereal2_cases[of b c]) auto lemma ereal_add_nonneg_eq_0_iff: fixes a b :: ereal shows "0 \ a \ 0 \ b \ a + b = 0 \ a = 0 \ b = 0" by (cases a b rule: ereal2_cases) auto lemma ereal_uminus_eq_reorder: "- a = b \ a = (-b::ereal)" by auto lemma ereal_uminus_less_reorder: "- a < b \ -b < (a::ereal)" by (subst (3) ereal_uminus_uminus[symmetric]) (simp only: ereal_minus_less_minus) lemma ereal_less_uminus_reorder: "a < - b \ b < - (a::ereal)" by (subst (3) ereal_uminus_uminus[symmetric]) (simp only: ereal_minus_less_minus) lemma ereal_uminus_le_reorder: "- a \ b \ -b \ (a::ereal)" by (subst (3) ereal_uminus_uminus[symmetric]) (simp only: ereal_minus_le_minus) lemmas ereal_uminus_reorder = ereal_uminus_eq_reorder ereal_uminus_less_reorder ereal_uminus_le_reorder lemma ereal_bot: fixes x :: ereal assumes "\B. x \ ereal B" shows "x = - \" proof (cases x) case (real r) with assms[of "r - 1"] show ?thesis by auto next case PInf with assms[of 0] show ?thesis by auto next case MInf then show ?thesis by simp qed lemma ereal_top: fixes x :: ereal assumes "\B. x \ ereal B" shows "x = \" proof (cases x) case (real r) with assms[of "r + 1"] show ?thesis by auto next case MInf with assms[of 0] show ?thesis by auto next case PInf then show ?thesis by simp qed lemma shows ereal_max[simp]: "ereal (max x y) = max (ereal x) (ereal y)" and ereal_min[simp]: "ereal (min x y) = min (ereal x) (ereal y)" by (simp_all add: min_def max_def) lemma ereal_max_0: "max 0 (ereal r) = ereal (max 0 r)" by (auto simp: zero_ereal_def) lemma fixes f :: "nat \ ereal" shows ereal_incseq_uminus[simp]: "incseq (\x. - f x) \ decseq f" and ereal_decseq_uminus[simp]: "decseq (\x. - f x) \ incseq f" unfolding decseq_def incseq_def by auto lemma incseq_ereal: "incseq f \ incseq (\x. ereal (f x))" unfolding incseq_def by auto lemma sum_ereal[simp]: "(\x\A. ereal (f x)) = ereal (\x\A. f x)" proof (cases "finite A") case True then show ?thesis by induct auto next case False then show ?thesis by simp qed lemma sum_list_ereal [simp]: "sum_list (map (\x. ereal (f x)) xs) = ereal (sum_list (map f xs))" by (induction xs) simp_all lemma sum_Pinfty: fixes f :: "'a \ ereal" shows "(\x\P. f x) = \ \ finite P \ (\i\P. f i = \)" proof safe assume *: "sum f P = \" show "finite P" proof (rule ccontr) assume "\ finite P" with * show False by auto qed show "\i\P. f i = \" proof (rule ccontr) assume "\ ?thesis" then have "\i. i \ P \ f i \ \" by auto with \finite P\ have "sum f P \ \" by induct auto with * show False by auto qed next fix i assume "finite P" and "i \ P" and "f i = \" then show "sum f P = \" proof induct case (insert x A) show ?case using insert by (cases "x = i") auto qed simp qed lemma sum_Inf: fixes f :: "'a \ ereal" shows "\sum f A\ = \ \ finite A \ (\i\A. \f i\ = \)" proof assume *: "\sum f A\ = \" have "finite A" by (rule ccontr) (insert *, auto) moreover have "\i\A. \f i\ = \" proof (rule ccontr) assume "\ ?thesis" then have "\i\A. \r. f i = ereal r" by auto from bchoice[OF this] obtain r where "\x\A. f x = ereal (r x)" .. with * show False by auto qed ultimately show "finite A \ (\i\A. \f i\ = \)" by auto next assume "finite A \ (\i\A. \f i\ = \)" then obtain i where "finite A" "i \ A" and "\f i\ = \" by auto then show "\sum f A\ = \" proof induct case (insert j A) then show ?case by (cases rule: ereal3_cases[of "f i" "f j" "sum f A"]) auto qed simp qed lemma sum_real_of_ereal: fixes f :: "'i \ ereal" assumes "\x. x \ S \ \f x\ \ \" shows "(\x\S. real_of_ereal (f x)) = real_of_ereal (sum f S)" proof - have "\x\S. \r. f x = ereal r" proof fix x assume "x \ S" from assms[OF this] show "\r. f x = ereal r" by (cases "f x") auto qed from bchoice[OF this] obtain r where "\x\S. f x = ereal (r x)" .. then show ?thesis by simp qed lemma sum_ereal_0: fixes f :: "'a \ ereal" assumes "finite A" and "\i. i \ A \ 0 \ f i" shows "(\x\A. f x) = 0 \ (\i\A. f i = 0)" proof assume "sum f A = 0" with assms show "\i\A. f i = 0" proof (induction A) case (insert a A) then have "f a = 0 \ (\a\A. f a) = 0" by (subst ereal_add_nonneg_eq_0_iff[symmetric]) (simp_all add: sum_nonneg) with insert show ?case by simp qed simp qed auto subsubsection "Multiplication" instantiation ereal :: "{comm_monoid_mult,sgn}" begin function sgn_ereal :: "ereal \ ereal" where "sgn (ereal r) = ereal (sgn r)" | "sgn (\::ereal) = 1" | "sgn (-\::ereal) = -1" by (auto intro: ereal_cases) termination by standard (rule wf_empty) function times_ereal where "ereal r * ereal p = ereal (r * p)" | "ereal r * \ = (if r = 0 then 0 else if r > 0 then \ else -\)" | "\ * ereal r = (if r = 0 then 0 else if r > 0 then \ else -\)" | "ereal r * -\ = (if r = 0 then 0 else if r > 0 then -\ else \)" | "-\ * ereal r = (if r = 0 then 0 else if r > 0 then -\ else \)" | "(\::ereal) * \ = \" | "-(\::ereal) * \ = -\" | "(\::ereal) * -\ = -\" | "-(\::ereal) * -\ = \" proof goal_cases case prems: (1 P x) then obtain a b where "x = (a, b)" by (cases x) auto with prems show P by (cases rule: ereal2_cases[of a b]) auto qed simp_all termination by (relation "{}") simp instance proof fix a b c :: ereal show "1 * a = a" by (cases a) (simp_all add: one_ereal_def) show "a * b = b * a" by (cases rule: ereal2_cases[of a b]) simp_all show "a * b * c = a * (b * c)" by (cases rule: ereal3_cases[of a b c]) (simp_all add: zero_ereal_def zero_less_mult_iff) qed end lemma [simp]: shows ereal_1_times: "ereal 1 * x = x" and times_ereal_1: "x * ereal 1 = x" by(simp_all flip: one_ereal_def) lemma one_not_le_zero_ereal[simp]: "\ (1 \ (0::ereal))" by (simp add: one_ereal_def zero_ereal_def) lemma real_ereal_1[simp]: "real_of_ereal (1::ereal) = 1" unfolding one_ereal_def by simp lemma real_of_ereal_le_1: fixes a :: ereal shows "a \ 1 \ real_of_ereal a \ 1" by (cases a) (auto simp: one_ereal_def) lemma abs_ereal_one[simp]: "\1\ = (1::ereal)" unfolding one_ereal_def by simp lemma ereal_mult_zero[simp]: fixes a :: ereal shows "a * 0 = 0" by (cases a) (simp_all add: zero_ereal_def) lemma ereal_zero_mult[simp]: fixes a :: ereal shows "0 * a = 0" by (cases a) (simp_all add: zero_ereal_def) lemma ereal_m1_less_0[simp]: "-(1::ereal) < 0" by (simp add: zero_ereal_def one_ereal_def) lemma ereal_times[simp]: "1 \ (\::ereal)" "(\::ereal) \ 1" "1 \ -(\::ereal)" "-(\::ereal) \ 1" by (auto simp: one_ereal_def) lemma ereal_plus_1[simp]: "1 + ereal r = ereal (r + 1)" "ereal r + 1 = ereal (r + 1)" "1 + -(\::ereal) = -\" "-(\::ereal) + 1 = -\" unfolding one_ereal_def by auto lemma ereal_zero_times[simp]: fixes a b :: ereal shows "a * b = 0 \ a = 0 \ b = 0" by (cases rule: ereal2_cases[of a b]) auto lemma ereal_mult_eq_PInfty[simp]: "a * b = (\::ereal) \ (a = \ \ b > 0) \ (a > 0 \ b = \) \ (a = -\ \ b < 0) \ (a < 0 \ b = -\)" by (cases rule: ereal2_cases[of a b]) auto lemma ereal_mult_eq_MInfty[simp]: "a * b = -(\::ereal) \ (a = \ \ b < 0) \ (a < 0 \ b = \) \ (a = -\ \ b > 0) \ (a > 0 \ b = -\)" by (cases rule: ereal2_cases[of a b]) auto lemma ereal_abs_mult: "\x * y :: ereal\ = \x\ * \y\" by (cases x y rule: ereal2_cases) (auto simp: abs_mult) lemma ereal_0_less_1[simp]: "0 < (1::ereal)" by (simp_all add: zero_ereal_def one_ereal_def) lemma ereal_mult_minus_left[simp]: fixes a b :: ereal shows "-a * b = - (a * b)" by (cases rule: ereal2_cases[of a b]) auto lemma ereal_mult_minus_right[simp]: fixes a b :: ereal shows "a * -b = - (a * b)" by (cases rule: ereal2_cases[of a b]) auto lemma ereal_mult_infty[simp]: "a * (\::ereal) = (if a = 0 then 0 else if 0 < a then \ else - \)" by (cases a) auto lemma ereal_infty_mult[simp]: "(\::ereal) * a = (if a = 0 then 0 else if 0 < a then \ else - \)" by (cases a) auto lemma ereal_mult_strict_right_mono: assumes "a < b" and "0 < c" and "c < (\::ereal)" shows "a * c < b * c" using assms by (cases rule: ereal3_cases[of a b c]) (auto simp: zero_le_mult_iff) lemma ereal_mult_strict_left_mono: "a < b \ 0 < c \ c < (\::ereal) \ c * a < c * b" using ereal_mult_strict_right_mono by (simp add: mult.commute[of c]) lemma ereal_mult_right_mono: fixes a b c :: ereal assumes "a \ b" "0 \ c" shows "a * c \ b * c" proof (cases "c = 0") case False with assms show ?thesis by (cases rule: ereal3_cases[of a b c]) auto qed auto lemma ereal_mult_left_mono: fixes a b c :: ereal shows "a \ b \ 0 \ c \ c * a \ c * b" using ereal_mult_right_mono by (simp add: mult.commute[of c]) lemma ereal_mult_mono: fixes a b c d::ereal assumes "b \ 0" "c \ 0" "a \ b" "c \ d" shows "a * c \ b * d" by (metis ereal_mult_right_mono mult.commute order_trans assms) lemma ereal_mult_mono': fixes a b c d::ereal assumes "a \ 0" "c \ 0" "a \ b" "c \ d" shows "a * c \ b * d" by (metis ereal_mult_right_mono mult.commute order_trans assms) lemma ereal_mult_mono_strict: fixes a b c d::ereal assumes "b > 0" "c > 0" "a < b" "c < d" shows "a * c < b * d" proof - have "c < \" using \c < d\ by auto then have "a * c < b * c" by (metis ereal_mult_strict_left_mono[OF assms(3) assms(2)] mult.commute) moreover have "b * c \ b * d" using assms(2) assms(4) by (simp add: assms(1) ereal_mult_left_mono less_imp_le) ultimately show ?thesis by simp qed lemma ereal_mult_mono_strict': fixes a b c d::ereal assumes "a > 0" "c > 0" "a < b" "c < d" shows "a * c < b * d" using assms ereal_mult_mono_strict by auto lemma zero_less_one_ereal[simp]: "0 \ (1::ereal)" by (simp add: one_ereal_def zero_ereal_def) lemma ereal_0_le_mult[simp]: "0 \ a \ 0 \ b \ 0 \ a * (b :: ereal)" by (cases rule: ereal2_cases[of a b]) auto lemma ereal_right_distrib: fixes r a b :: ereal shows "0 \ a \ 0 \ b \ r * (a + b) = r * a + r * b" by (cases rule: ereal3_cases[of r a b]) (simp_all add: field_simps) lemma ereal_left_distrib: fixes r a b :: ereal shows "0 \ a \ 0 \ b \ (a + b) * r = a * r + b * r" by (cases rule: ereal3_cases[of r a b]) (simp_all add: field_simps) lemma ereal_mult_le_0_iff: fixes a b :: ereal shows "a * b \ 0 \ (0 \ a \ b \ 0) \ (a \ 0 \ 0 \ b)" by (cases rule: ereal2_cases[of a b]) (simp_all add: mult_le_0_iff) lemma ereal_zero_le_0_iff: fixes a b :: ereal shows "0 \ a * b \ (0 \ a \ 0 \ b) \ (a \ 0 \ b \ 0)" by (cases rule: ereal2_cases[of a b]) (simp_all add: zero_le_mult_iff) lemma ereal_mult_less_0_iff: fixes a b :: ereal shows "a * b < 0 \ (0 < a \ b < 0) \ (a < 0 \ 0 < b)" by (cases rule: ereal2_cases[of a b]) (simp_all add: mult_less_0_iff) lemma ereal_zero_less_0_iff: fixes a b :: ereal shows "0 < a * b \ (0 < a \ 0 < b) \ (a < 0 \ b < 0)" by (cases rule: ereal2_cases[of a b]) (simp_all add: zero_less_mult_iff) lemma ereal_left_mult_cong: fixes a b c :: ereal shows "c = d \ (d \ 0 \ a = b) \ a * c = b * d" by (cases "c = 0") simp_all lemma ereal_right_mult_cong: fixes a b c :: ereal shows "c = d \ (d \ 0 \ a = b) \ c * a = d * b" by (cases "c = 0") simp_all lemma ereal_distrib: fixes a b c :: ereal assumes "a \ \ \ b \ -\" and "a \ -\ \ b \ \" and "\c\ \ \" shows "(a + b) * c = a * c + b * c" using assms by (cases rule: ereal3_cases[of a b c]) (simp_all add: field_simps) lemma numeral_eq_ereal [simp]: "numeral w = ereal (numeral w)" proof (induct w rule: num_induct) case One then show ?case by simp next case (inc x) then show ?case by (simp add: inc numeral_inc) qed lemma distrib_left_ereal_nn: "c \ 0 \ (x + y) * ereal c = x * ereal c + y * ereal c" by(cases x y rule: ereal2_cases)(simp_all add: ring_distribs) lemma sum_ereal_right_distrib: fixes f :: "'a \ ereal" shows "(\i. i \ A \ 0 \ f i) \ r * sum f A = (\n\A. r * f n)" by (induct A rule: infinite_finite_induct) (auto simp: ereal_right_distrib sum_nonneg) lemma sum_ereal_left_distrib: "(\i. i \ A \ 0 \ f i) \ sum f A * r = (\n\A. f n * r :: ereal)" using sum_ereal_right_distrib[of A f r] by (simp add: mult_ac) lemma sum_distrib_right_ereal: "c \ 0 \ sum f A * ereal c = (\x\A. f x * c :: ereal)" by(subst sum_comp_morphism[where h="\x. x * ereal c", symmetric])(simp_all add: distrib_left_ereal_nn) lemma ereal_le_epsilon: fixes x y :: ereal assumes "\e. 0 < e \ x \ y + e" shows "x \ y" proof (cases "x = -\ \ x = \ \ y = -\ \ y = \") case True then show ?thesis using assms[of 1] by auto next case False then obtain p q where "x = ereal p" "y = ereal q" by (metis MInfty_eq_minfinity ereal.distinct(3) uminus_ereal.elims) then show ?thesis by (metis assms field_le_epsilon ereal_less(2) ereal_less_eq(3) plus_ereal.simps(1)) qed lemma ereal_le_epsilon2: fixes x y :: ereal assumes "\e::real. 0 < e \ x \ y + ereal e" shows "x \ y" proof (rule ereal_le_epsilon) show "\\::ereal. 0 < \ \ x \ y + \" using assms less_ereal.elims(2) zero_less_real_of_ereal by fastforce qed lemma ereal_le_real: fixes x y :: ereal assumes "\z. x \ ereal z \ y \ ereal z" shows "y \ x" by (metis assms ereal_bot ereal_cases ereal_infty_less_eq(2) ereal_less_eq(1) linorder_le_cases) lemma prod_ereal_0: fixes f :: "'a \ ereal" shows "(\i\A. f i) = 0 \ finite A \ (\i\A. f i = 0)" proof (cases "finite A") case True then show ?thesis by (induct A) auto qed auto lemma prod_ereal_pos: fixes f :: "'a \ ereal" assumes pos: "\i. i \ I \ 0 \ f i" shows "0 \ (\i\I. f i)" proof (cases "finite I") case True from this pos show ?thesis by induct auto qed auto lemma prod_PInf: fixes f :: "'a \ ereal" assumes "\i. i \ I \ 0 \ f i" shows "(\i\I. f i) = \ \ finite I \ (\i\I. f i = \) \ (\i\I. f i \ 0)" proof (cases "finite I") case True from this assms show ?thesis proof (induct I) case (insert i I) then have pos: "0 \ f i" "0 \ prod f I" by (auto intro!: prod_ereal_pos) from insert have "(\j\insert i I. f j) = \ \ prod f I * f i = \" by auto also have "\ \ (prod f I = \ \ f i = \) \ f i \ 0 \ prod f I \ 0" using prod_ereal_pos[of I f] pos by (cases rule: ereal2_cases[of "f i" "prod f I"]) auto also have "\ \ finite (insert i I) \ (\j\insert i I. f j = \) \ (\j\insert i I. f j \ 0)" using insert by (auto simp: prod_ereal_0) finally show ?case . qed simp qed auto lemma prod_ereal: "(\i\A. ereal (f i)) = ereal (prod f A)" proof (cases "finite A") case True then show ?thesis by induct (auto simp: one_ereal_def) next case False then show ?thesis by (simp add: one_ereal_def) qed subsubsection \Power\ lemma ereal_power[simp]: "(ereal x) ^ n = ereal (x^n)" by (induct n) (auto simp: one_ereal_def) lemma ereal_power_PInf[simp]: "(\::ereal) ^ n = (if n = 0 then 1 else \)" by (induct n) (auto simp: one_ereal_def) lemma ereal_power_uminus[simp]: fixes x :: ereal shows "(- x) ^ n = (if even n then x ^ n else - (x^n))" by (induct n) (auto simp: one_ereal_def) lemma ereal_power_numeral[simp]: "(numeral num :: ereal) ^ n = ereal (numeral num ^ n)" by (induct n) (auto simp: one_ereal_def) lemma zero_le_power_ereal[simp]: fixes a :: ereal assumes "0 \ a" shows "0 \ a ^ n" using assms by (induct n) (auto simp: ereal_zero_le_0_iff) subsubsection \Subtraction\ lemma ereal_minus_minus_image[simp]: fixes S :: "ereal set" shows "uminus ` uminus ` S = S" by (auto simp: image_iff) lemma ereal_uminus_lessThan[simp]: fixes a :: ereal shows "uminus ` {.. - ereal r = -\" "ereal r - \ = -\" "(\::ereal) - x = \" "-(\::ereal) - \ = -\" "x - -y = x + y" "x - 0 = x" "0 - x = -x" by (simp_all add: minus_ereal_def) lemma ereal_x_minus_x[simp]: "x - x = (if \x\ = \ then \ else 0::ereal)" by (cases x) simp_all lemma ereal_eq_minus_iff: fixes x y z :: ereal shows "x = z - y \ (\y\ \ \ \ x + y = z) \ (y = -\ \ x = \) \ (y = \ \ z = \ \ x = \) \ (y = \ \ z \ \ \ x = -\)" by (cases rule: ereal3_cases[of x y z]) auto lemma ereal_eq_minus: fixes x y z :: ereal shows "\y\ \ \ \ x = z - y \ x + y = z" by (auto simp: ereal_eq_minus_iff) lemma ereal_less_minus_iff: fixes x y z :: ereal shows "x < z - y \ (y = \ \ z = \ \ x \ \) \ (y = -\ \ x \ \) \ (\y\ \ \\ x + y < z)" by (cases rule: ereal3_cases[of x y z]) auto lemma ereal_less_minus: fixes x y z :: ereal shows "\y\ \ \ \ x < z - y \ x + y < z" by (auto simp: ereal_less_minus_iff) lemma ereal_le_minus_iff: fixes x y z :: ereal shows "x \ z - y \ (y = \ \ z \ \ \ x = -\) \ (\y\ \ \ \ x + y \ z)" by (cases rule: ereal3_cases[of x y z]) auto lemma ereal_le_minus: fixes x y z :: ereal shows "\y\ \ \ \ x \ z - y \ x + y \ z" by (auto simp: ereal_le_minus_iff) lemma ereal_minus_less_iff: fixes x y z :: ereal shows "x - y < z \ y \ -\ \ (y = \ \ x \ \ \ z \ -\) \ (y \ \ \ x < z + y)" by (cases rule: ereal3_cases[of x y z]) auto lemma ereal_minus_less: fixes x y z :: ereal shows "\y\ \ \ \ x - y < z \ x < z + y" by (auto simp: ereal_minus_less_iff) lemma ereal_minus_le_iff: fixes x y z :: ereal shows "x - y \ z \ (y = -\ \ z = \) \ (y = \ \ x = \ \ z = \) \ (\y\ \ \ \ x \ z + y)" by (cases rule: ereal3_cases[of x y z]) auto lemma ereal_minus_le: fixes x y z :: ereal shows "\y\ \ \ \ x - y \ z \ x \ z + y" by (auto simp: ereal_minus_le_iff) lemma ereal_minus_eq_minus_iff: fixes a b c :: ereal shows "a - b = a - c \ b = c \ a = \ \ (a = -\ \ b \ -\ \ c \ -\)" by (cases rule: ereal3_cases[of a b c]) auto lemma ereal_add_le_add_iff: fixes a b c :: ereal shows "c + a \ c + b \ a \ b \ c = \ \ (c = -\ \ a \ \ \ b \ \)" by (cases rule: ereal3_cases[of a b c]) (simp_all add: field_simps) lemma ereal_add_le_add_iff2: fixes a b c :: ereal shows "a + c \ b + c \ a \ b \ c = \ \ (c = -\ \ a \ \ \ b \ \)" by(cases rule: ereal3_cases[of a b c])(simp_all add: field_simps) lemma ereal_mult_le_mult_iff: fixes a b c :: ereal shows "\c\ \ \ \ c * a \ c * b \ (0 < c \ a \ b) \ (c < 0 \ b \ a)" by (cases rule: ereal3_cases[of a b c]) (simp_all add: mult_le_cancel_left) lemma ereal_minus_mono: fixes A B C D :: ereal assumes "A \ B" "D \ C" shows "A - C \ B - D" using assms by (cases rule: ereal3_cases[case_product ereal_cases, of A B C D]) simp_all lemma ereal_mono_minus_cancel: fixes a b c :: ereal shows "c - a \ c - b \ 0 \ c \ c < \ \ b \ a" by (cases a b c rule: ereal3_cases) auto lemma real_of_ereal_minus: fixes a b :: ereal shows "real_of_ereal (a - b) = (if \a\ = \ \ \b\ = \ then 0 else real_of_ereal a - real_of_ereal b)" by (cases rule: ereal2_cases[of a b]) auto lemma real_of_ereal_minus': "\x\ = \ \ \y\ = \ \ real_of_ereal x - real_of_ereal y = real_of_ereal (x - y :: ereal)" by(subst real_of_ereal_minus) auto lemma ereal_diff_positive: fixes a b :: ereal shows "a \ b \ 0 \ b - a" by (cases rule: ereal2_cases[of a b]) auto lemma ereal_between: fixes x e :: ereal assumes "\x\ \ \" and "0 < e" shows "x - e < x" and "x < x + e" using assms by (cases x, cases e, auto)+ lemma ereal_minus_eq_PInfty_iff: fixes x y :: ereal shows "x - y = \ \ y = -\ \ x = \" by (cases x y rule: ereal2_cases) simp_all lemma ereal_diff_add_eq_diff_diff_swap: fixes x y z :: ereal shows "\y\ \ \ \ x - (y + z) = x - y - z" by(cases x y z rule: ereal3_cases) simp_all lemma ereal_diff_add_assoc2: fixes x y z :: ereal shows "x + y - z = x - z + y" by(cases x y z rule: ereal3_cases) simp_all lemma ereal_add_uminus_conv_diff: fixes x y z :: ereal shows "- x + y = y - x" by(cases x y rule: ereal2_cases) simp_all lemma ereal_minus_diff_eq: fixes x y :: ereal shows "\ x = \ \ y \ \; x = -\ \ y \ - \ \ \ - (x - y) = y - x" by(cases x y rule: ereal2_cases) simp_all lemma ediff_le_self [simp]: "x - y \ (x :: enat)" by(cases x y rule: enat.exhaust[case_product enat.exhaust]) simp_all lemma ereal_abs_diff: fixes a b::ereal shows "abs(a-b) \ abs a + abs b" by (cases rule: ereal2_cases[of a b]) (auto) subsubsection \Division\ instantiation ereal :: inverse begin function inverse_ereal where "inverse (ereal r) = (if r = 0 then \ else ereal (inverse r))" | "inverse (\::ereal) = 0" | "inverse (-\::ereal) = 0" by (auto intro: ereal_cases) termination by (relation "{}") simp definition "x div y = x * inverse (y :: ereal)" instance .. end lemma real_of_ereal_inverse[simp]: fixes a :: ereal shows "real_of_ereal (inverse a) = 1 / real_of_ereal a" by (cases a) (auto simp: inverse_eq_divide) lemma ereal_inverse[simp]: "inverse (0::ereal) = \" "inverse (1::ereal) = 1" by (simp_all add: one_ereal_def zero_ereal_def) lemma ereal_divide[simp]: "ereal r / ereal p = (if p = 0 then ereal r * \ else ereal (r / p))" unfolding divide_ereal_def by (auto simp: divide_real_def) lemma ereal_divide_same[simp]: fixes x :: ereal shows "x / x = (if \x\ = \ \ x = 0 then 0 else 1)" by (cases x) (simp_all add: divide_real_def divide_ereal_def one_ereal_def) lemma ereal_inv_inv[simp]: fixes x :: ereal shows "inverse (inverse x) = (if x \ -\ then x else \)" by (cases x) auto lemma ereal_inverse_minus[simp]: fixes x :: ereal shows "inverse (- x) = (if x = 0 then \ else -inverse x)" by (cases x) simp_all lemma ereal_uminus_divide[simp]: fixes x y :: ereal shows "- x / y = - (x / y)" unfolding divide_ereal_def by simp lemma ereal_divide_Infty[simp]: fixes x :: ereal shows "x / \ = 0" "x / -\ = 0" unfolding divide_ereal_def by simp_all lemma ereal_divide_one[simp]: "x / 1 = (x::ereal)" unfolding divide_ereal_def by simp lemma ereal_divide_ereal[simp]: "\ / ereal r = (if 0 \ r then \ else -\)" unfolding divide_ereal_def by simp lemma ereal_inverse_nonneg_iff: "0 \ inverse (x :: ereal) \ 0 \ x \ x = -\" by (cases x) auto lemma inverse_ereal_ge0I: "0 \ (x :: ereal) \ 0 \ inverse x" by(cases x) simp_all lemma zero_le_divide_ereal[simp]: fixes a :: ereal assumes "0 \ a" and "0 \ b" shows "0 \ a / b" using assms by (cases rule: ereal2_cases[of a b]) (auto simp: zero_le_divide_iff) lemma ereal_le_divide_pos: fixes x y z :: ereal shows "x > 0 \ x \ \ \ y \ z / x \ x * y \ z" by (cases rule: ereal3_cases[of x y z]) (auto simp: field_simps) lemma ereal_divide_le_pos: fixes x y z :: ereal shows "x > 0 \ x \ \ \ z / x \ y \ z \ x * y" by (cases rule: ereal3_cases[of x y z]) (auto simp: field_simps) lemma ereal_le_divide_neg: fixes x y z :: ereal shows "x < 0 \ x \ -\ \ y \ z / x \ z \ x * y" by (cases rule: ereal3_cases[of x y z]) (auto simp: field_simps) lemma ereal_divide_le_neg: fixes x y z :: ereal shows "x < 0 \ x \ -\ \ z / x \ y \ x * y \ z" by (cases rule: ereal3_cases[of x y z]) (auto simp: field_simps) lemma ereal_inverse_antimono_strict: fixes x y :: ereal shows "0 \ x \ x < y \ inverse y < inverse x" by (cases rule: ereal2_cases[of x y]) auto lemma ereal_inverse_antimono: fixes x y :: ereal shows "0 \ x \ x \ y \ inverse y \ inverse x" by (cases rule: ereal2_cases[of x y]) auto lemma inverse_inverse_Pinfty_iff[simp]: fixes x :: ereal shows "inverse x = \ \ x = 0" by (cases x) auto lemma ereal_inverse_eq_0: fixes x :: ereal shows "inverse x = 0 \ x = \ \ x = -\" by (cases x) auto lemma ereal_0_gt_inverse: fixes x :: ereal shows "0 < inverse x \ x \ \ \ 0 \ x" by (cases x) auto lemma ereal_inverse_le_0_iff: fixes x :: ereal shows "inverse x \ 0 \ x < 0 \ x = \" by(cases x) auto lemma ereal_divide_eq_0_iff: "x / y = 0 \ x = 0 \ \y :: ereal\ = \" by(cases x y rule: ereal2_cases) simp_all lemma ereal_mult_less_right: fixes a b c :: ereal assumes "b * a < c * a" and "0 < a" and "a < \" shows "b < c" using assms by (cases rule: ereal3_cases[of a b c]) (auto split: if_split_asm simp: zero_less_mult_iff zero_le_mult_iff) lemma ereal_mult_divide: fixes a b :: ereal shows "0 < b \ b < \ \ b * (a / b) = a" by (cases a b rule: ereal2_cases) auto lemma ereal_power_divide: fixes x y :: ereal shows "y \ 0 \ (x / y) ^ n = x^n / y^n" by (cases rule: ereal2_cases [of x y]) (auto simp: one_ereal_def zero_ereal_def power_divide zero_le_power_eq) lemma ereal_le_mult_one_interval: fixes x y :: ereal assumes y: "y \ -\" assumes z: "\z. 0 < z \ z < 1 \ z * x \ y" shows "x \ y" proof (cases x) case PInf with z[of "1 / 2"] show "x \ y" by (simp add: one_ereal_def) next case (real r) note r = this show "x \ y" proof (cases y) case (real p) note p = this have "r \ p" proof (rule field_le_mult_one_interval) fix z :: real assume "0 < z" and "z < 1" with z[of "ereal z"] show "z * r \ p" using p r by (auto simp: zero_le_mult_iff one_ereal_def) qed then show "x \ y" using p r by simp qed (insert y, simp_all) qed simp lemma ereal_divide_right_mono[simp]: fixes x y z :: ereal assumes "x \ y" and "0 < z" shows "x / z \ y / z" using assms by (cases x y z rule: ereal3_cases) (auto intro: divide_right_mono) lemma ereal_divide_left_mono[simp]: fixes x y z :: ereal assumes "y \ x" and "0 < z" and "0 < x * y" shows "z / x \ z / y" using assms by (cases x y z rule: ereal3_cases) (auto intro: divide_left_mono simp: field_simps zero_less_mult_iff mult_less_0_iff split: if_split_asm) lemma ereal_divide_zero_left[simp]: fixes a :: ereal shows "0 / a = 0" by (cases a) (auto simp: zero_ereal_def) lemma ereal_times_divide_eq_left[simp]: fixes a b c :: ereal shows "b / c * a = b * a / c" by (cases a b c rule: ereal3_cases) (auto simp: field_simps zero_less_mult_iff mult_less_0_iff) lemma ereal_times_divide_eq: "a * (b / c :: ereal) = a * b / c" by (cases a b c rule: ereal3_cases) (auto simp: field_simps zero_less_mult_iff) lemma ereal_inverse_real [simp]: "\z\ \ \ \ z \ 0 \ ereal (inverse (real_of_ereal z)) = inverse z" by auto lemma ereal_inverse_mult: "a \ 0 \ b \ 0 \ inverse (a * (b::ereal)) = inverse a * inverse b" by (cases a; cases b) auto lemma inverse_eq_infinity_iff_eq_zero [simp]: "1/(x::ereal) = \ \ x = 0" by (simp add: divide_ereal_def) lemma ereal_distrib_left: fixes a b c :: ereal assumes "a \ \ \ b \ -\" and "a \ -\ \ b \ \" and "\c\ \ \" shows "c * (a + b) = c * a + c * b" using assms by (cases rule: ereal3_cases[of a b c]) (simp_all add: field_simps) lemma ereal_distrib_minus_left: fixes a b c :: ereal assumes "a \ \ \ b \ \" and "a \ -\ \ b \ -\" and "\c\ \ \" shows "c * (a - b) = c * a - c * b" using assms by (cases rule: ereal3_cases[of a b c]) (simp_all add: field_simps) lemma ereal_distrib_minus_right: fixes a b c :: ereal assumes "a \ \ \ b \ \" and "a \ -\ \ b \ -\" and "\c\ \ \" shows "(a - b) * c = a * c - b * c" using assms by (cases rule: ereal3_cases[of a b c]) (simp_all add: field_simps) subsection "Complete lattice" instantiation ereal :: lattice begin definition [simp]: "sup x y = (max x y :: ereal)" definition [simp]: "inf x y = (min x y :: ereal)" instance by standard simp_all end instantiation ereal :: complete_lattice begin definition "bot = (-\::ereal)" definition "top = (\::ereal)" definition "Sup S = (SOME x :: ereal. (\y\S. y \ x) \ (\z. (\y\S. y \ z) \ x \ z))" definition "Inf S = (SOME x :: ereal. (\y\S. x \ y) \ (\z. (\y\S. z \ y) \ z \ x))" lemma ereal_complete_Sup: fixes S :: "ereal set" shows "\x. (\y\S. y \ x) \ (\z. (\y\S. y \ z) \ x \ z)" proof (cases "\x. \a\S. a \ ereal x") case True then obtain y where y: "a \ ereal y" if "a\S" for a by auto then have "\ \ S" by force show ?thesis proof (cases "S \ {-\} \ S \ {}") case True with \\ \ S\ obtain x where x: "x \ S" "\x\ \ \" by auto obtain s where s: "\x\ereal -` S. x \ s" "(\x\ereal -` S. x \ z) \ s \ z" for z proof (atomize_elim, rule complete_real) show "\x. x \ ereal -` S" using x by auto show "\z. \x\ereal -` S. x \ z" by (auto dest: y intro!: exI[of _ y]) qed show ?thesis proof (safe intro!: exI[of _ "ereal s"]) fix y assume "y \ S" with s \\ \ S\ show "y \ ereal s" by (cases y) auto next fix z assume "\y\S. y \ z" with \S \ {-\} \ S \ {}\ show "ereal s \ z" by (cases z) (auto intro!: s) qed next case False then show ?thesis by (auto intro!: exI[of _ "-\"]) qed next case False then show ?thesis by (fastforce intro!: exI[of _ \] ereal_top intro: order_trans dest: less_imp_le simp: not_le) qed lemma ereal_complete_uminus_eq: fixes S :: "ereal set" shows "(\y\uminus`S. y \ x) \ (\z. (\y\uminus`S. y \ z) \ x \ z) \ (\y\S. -x \ y) \ (\z. (\y\S. z \ y) \ z \ -x)" by simp (metis ereal_minus_le_minus ereal_uminus_uminus) lemma ereal_complete_Inf: "\x. (\y\S::ereal set. x \ y) \ (\z. (\y\S. z \ y) \ z \ x)" using ereal_complete_Sup[of "uminus ` S"] unfolding ereal_complete_uminus_eq by auto instance proof show "Sup {} = (bot::ereal)" using ereal_bot by (auto simp: bot_ereal_def Sup_ereal_def) show "Inf {} = (top::ereal)" unfolding top_ereal_def Inf_ereal_def using ereal_infty_less_eq(1) ereal_less_eq(1) by blast qed (auto intro: someI2_ex ereal_complete_Sup ereal_complete_Inf simp: Sup_ereal_def Inf_ereal_def bot_ereal_def top_ereal_def) end instance ereal :: complete_linorder .. instance ereal :: linear_continuum proof show "\a b::ereal. a \ b" using zero_neq_one by blast qed lemma min_PInf [simp]: "min (\::ereal) x = x" by (metis min_top top_ereal_def) lemma min_PInf2 [simp]: "min x (\::ereal) = x" by (metis min_top2 top_ereal_def) lemma max_PInf [simp]: "max (\::ereal) x = \" by (metis max_top top_ereal_def) lemma max_PInf2 [simp]: "max x (\::ereal) = \" by (metis max_top2 top_ereal_def) lemma min_MInf [simp]: "min (-\::ereal) x = -\" by (metis min_bot bot_ereal_def) lemma min_MInf2 [simp]: "min x (-\::ereal) = -\" by (metis min_bot2 bot_ereal_def) lemma max_MInf [simp]: "max (-\::ereal) x = x" by (metis max_bot bot_ereal_def) lemma max_MInf2 [simp]: "max x (-\::ereal) = x" by (metis max_bot2 bot_ereal_def) subsection \Extended real intervals\ lemma real_greaterThanLessThan_infinity_eq: "real_of_ereal ` {N::ereal<..<\} = (if N = \ then {} else if N = -\ then UNIV else {real_of_ereal N<..})" by (force simp: real_less_ereal_iff intro!: image_eqI[where x="ereal _"] elim!: less_ereal.elims) lemma real_greaterThanLessThan_minus_infinity_eq: "real_of_ereal ` {-\<.. then UNIV else if N = -\ then {} else {..<..}" by (auto simp: ereal_uminus_less_reorder intro!: image_eqI[where x="-x" for x]) also note real_greaterThanLessThan_infinity_eq finally show ?thesis by (auto intro!: image_eqI[where x="-x" for x]) qed lemma real_greaterThanLessThan_inter: "real_of_ereal ` {N<..<.. real_of_ereal ` {N<..<\}" by (force elim!: less_ereal.elims) lemma real_atLeastGreaterThan_eq: "real_of_ereal ` {N<.. then {} else if N = -\ then (if M = \ then UNIV else if M = -\ then {} else {..< real_of_ereal M}) else if M = - \ then {} else if M = \ then {real_of_ereal N<..} else {real_of_ereal N <..< real_of_ereal M})" proof (cases "M = -\ \ M = \ \ N = -\ \ N = \") case True then show ?thesis by (auto simp: real_greaterThanLessThan_minus_infinity_eq real_greaterThanLessThan_infinity_eq ) next case False then obtain p q where "M = ereal p" "N = ereal q" by (metis MInfty_eq_minfinity ereal.distinct(3) uminus_ereal.elims) moreover have "\x. \q < x; x < p\ \ x \ real_of_ereal ` {ereal q<.. then if b = \ then UNIV else {.. then {real_of_ereal a<..} else {real_of_ereal a <..< real_of_ereal b}) else {})" by (cases a; cases b; simp add: real_atLeastGreaterThan_eq not_less) lemma fixes a b c::ereal shows not_inftyI: "a < b \ b < c \ abs b \ \" by force lemma interval_neqs: fixes r s t::real shows "{r<.. {t<..}" and "{r<.. {.. UNIV" and "{r<..} \ {.. UNIV" and "{.. UNIV" and "{} \ {r<..}" and "{} \ {.. s \ u \ t \ r = t \ s = u)" by (metis cInf_greaterThanLessThan cSup_greaterThanLessThan greaterThanLessThan_empty_iff not_le) lemma real_of_ereal_image_greaterThanLessThan_iff: "real_of_ereal ` {a <..< b} = real_of_ereal ` {c <..< d} \ (a \ b \ c \ d \ a = c \ b = d)" unfolding real_atLeastGreaterThan_eq by (cases a; cases b; cases c; cases d; simp add: greaterThanLessThan_eq_iff interval_neqs interval_neqs[symmetric]) lemma uminus_image_real_of_ereal_image_greaterThanLessThan: "uminus ` real_of_ereal ` {l <..< u} = real_of_ereal ` {-u <..< -l}" by (force simp: algebra_simps ereal_less_uminus_reorder ereal_uminus_less_reorder intro: image_eqI[where x="-x" for x]) lemma add_image_real_of_ereal_image_greaterThanLessThan: "(+) c ` real_of_ereal ` {l <..< u} = real_of_ereal ` {c + l <..< c + u}" apply safe subgoal for x using ereal_less_add[of c] by (force simp: real_of_ereal_add add.commute) subgoal for _ x by (force simp: add.commute real_of_ereal_minus ereal_minus_less ereal_less_minus intro: image_eqI[where x="x - c"]) done lemma add2_image_real_of_ereal_image_greaterThanLessThan: "(\x. x + c) ` real_of_ereal ` {l <..< u} = real_of_ereal ` {l + c <..< u + c}" using add_image_real_of_ereal_image_greaterThanLessThan[of c l u] by (metis add.commute image_cong) lemma minus_image_real_of_ereal_image_greaterThanLessThan: "(-) c ` real_of_ereal ` {l <..< u} = real_of_ereal ` {c - u <..< c - l}" (is "?l = ?r") proof - have "?l = (+) c ` uminus ` real_of_ereal ` {l <..< u}" by auto also note uminus_image_real_of_ereal_image_greaterThanLessThan also note add_image_real_of_ereal_image_greaterThanLessThan finally show ?thesis by (simp add: minus_ereal_def) qed lemma real_ereal_bound_lemma_up: assumes "s \ real_of_ereal ` {a<.. real_of_ereal ` {a<.. t" shows "b \ \" proof (cases b) case PInf then show ?thesis using assms apply clarsimp by (metis UNIV_I assms(1) ereal_less_PInfty greaterThan_iff less_eq_ereal_def less_le_trans real_image_ereal_ivl) qed auto lemma real_ereal_bound_lemma_down: assumes s: "s \ real_of_ereal ` {a<.. real_of_ereal ` {a<.. s" shows "a \ - \" proof (cases b) case (real r) then show ?thesis using assms real_greaterThanLessThan_minus_infinity_eq by force next case PInf then show ?thesis using t real_greaterThanLessThan_infinity_eq by auto next case MInf then show ?thesis using s by auto qed subsection "Topological space" instantiation ereal :: linear_continuum_topology begin definition "open_ereal" :: "ereal set \ bool" where open_ereal_generated: "open_ereal = generate_topology (range lessThan \ range greaterThan)" instance by standard (simp add: open_ereal_generated) end lemma continuous_on_ereal[continuous_intros]: assumes f: "continuous_on s f" shows "continuous_on s (\x. ereal (f x))" by (rule continuous_on_compose2 [OF continuous_onI_mono[of ereal UNIV] f]) auto lemma tendsto_ereal[tendsto_intros, simp, intro]: "(f \ x) F \ ((\x. ereal (f x)) \ ereal x) F" using isCont_tendsto_compose[of x ereal f F] continuous_on_ereal[of UNIV "\x. x"] by (simp add: continuous_on_eq_continuous_at) lemma tendsto_uminus_ereal[tendsto_intros, simp, intro]: assumes "(f \ x) F" shows "((\x. - f x::ereal) \ - x) F" proof (rule tendsto_compose[OF order_tendstoI assms]) show "\a. a < - x \ \\<^sub>F x in at x. a < - x" by (metis ereal_less_uminus_reorder eventually_at_topological lessThan_iff open_lessThan) show "\a. - x < a \ \\<^sub>F x in at x. - x < a" by (metis ereal_uminus_reorder(2) eventually_at_topological greaterThan_iff open_greaterThan) qed lemma at_infty_ereal_eq_at_top: "at \ = filtermap ereal at_top" unfolding filter_eq_iff eventually_at_filter eventually_at_top_linorder eventually_filtermap top_ereal_def[symmetric] apply (subst eventually_nhds_top[of 0]) apply (auto simp: top_ereal_def less_le ereal_all_split ereal_ex_split) apply (metis PInfty_neq_ereal(2) ereal_less_eq(3) ereal_top le_cases order_trans) done lemma ereal_Lim_uminus: "(f \ f0) net \ ((\x. - f x::ereal) \ - f0) net" using tendsto_uminus_ereal[of f f0 net] tendsto_uminus_ereal[of "\x. - f x" "- f0" net] by auto lemma ereal_divide_less_iff: "0 < (c::ereal) \ c < \ \ a / c < b \ a < b * c" by (cases a b c rule: ereal3_cases) (auto simp: field_simps) lemma ereal_less_divide_iff: "0 < (c::ereal) \ c < \ \ a < b / c \ a * c < b" by (cases a b c rule: ereal3_cases) (auto simp: field_simps) lemma tendsto_cmult_ereal[tendsto_intros, simp, intro]: assumes c: "\c\ \ \" and f: "(f \ x) F" shows "((\x. c * f x::ereal) \ c * x) F" proof - { fix c :: ereal assume "0 < c" "c < \" then have "((\x. c * f x::ereal) \ c * x) F" apply (intro tendsto_compose[OF _ f]) apply (auto intro!: order_tendstoI simp: eventually_at_topological) apply (rule_tac x="{a/c <..}" in exI) apply (auto split: ereal.split simp: ereal_divide_less_iff mult.commute) [] apply (rule_tac x="{..< a/c}" in exI) apply (auto split: ereal.split simp: ereal_less_divide_iff mult.commute) [] done } note * = this have "((0 < c \ c < \) \ (-\ < c \ c < 0) \ c = 0)" using c by (cases c) auto then show ?thesis proof (elim disjE conjE) assume "- \ < c" "c < 0" then have "0 < - c" "- c < \" by (auto simp: ereal_uminus_reorder ereal_less_uminus_reorder[of 0]) then have "((\x. (- c) * f x) \ (- c) * x) F" by (rule *) from tendsto_uminus_ereal[OF this] show ?thesis by simp qed (auto intro!: *) qed lemma tendsto_cmult_ereal_not_0[tendsto_intros, simp, intro]: assumes "x \ 0" and f: "(f \ x) F" shows "((\x. c * f x::ereal) \ c * x) F" proof cases assume "\c\ = \" show ?thesis proof (rule filterlim_cong[THEN iffD1, OF refl refl _ tendsto_const]) have "0 < x \ x < 0" using \x \ 0\ by (auto simp add: neq_iff) then show "eventually (\x'. c * x = c * f x') F" proof assume "0 < x" from order_tendstoD(1)[OF f this] show ?thesis by eventually_elim (insert \0 \\c\ = \\, auto) next assume "x < 0" from order_tendstoD(2)[OF f this] show ?thesis by eventually_elim (insert \x<0\ \\c\ = \\, auto) qed qed qed (rule tendsto_cmult_ereal[OF _ f]) lemma tendsto_cadd_ereal[tendsto_intros, simp, intro]: assumes c: "y \ - \" "x \ - \" and f: "(f \ x) F" shows "((\x. f x + y::ereal) \ x + y) F" apply (intro tendsto_compose[OF _ f]) apply (auto intro!: order_tendstoI simp: eventually_at_topological) apply (rule_tac x="{a - y <..}" in exI) apply (auto split: ereal.split simp: ereal_minus_less_iff c) [] apply (rule_tac x="{..< a - y}" in exI) apply (auto split: ereal.split simp: ereal_less_minus_iff c) [] done lemma tendsto_add_left_ereal[tendsto_intros, simp, intro]: assumes c: "\y\ \ \" and f: "(f \ x) F" shows "((\x. f x + y::ereal) \ x + y) F" apply (intro tendsto_compose[OF _ f]) apply (auto intro!: order_tendstoI simp: eventually_at_topological) apply (rule_tac x="{a - y <..}" in exI) apply (insert c, auto split: ereal.split simp: ereal_minus_less_iff) [] apply (rule_tac x="{..< a - y}" in exI) apply (auto split: ereal.split simp: ereal_less_minus_iff c) [] done lemma continuous_at_ereal[continuous_intros]: "continuous F f \ continuous F (\x. ereal (f x))" unfolding continuous_def by auto lemma ereal_Sup: assumes *: "\SUP a\A. ereal a\ \ \" shows "ereal (Sup A) = (SUP a\A. ereal a)" proof (rule continuous_at_Sup_mono) obtain r where r: "ereal r = (SUP a\A. ereal a)" "A \ {}" using * by (force simp: bot_ereal_def) then show "bdd_above A" "A \ {}" by (auto intro!: SUP_upper bdd_aboveI[of _ r] simp flip: ereal_less_eq) qed (auto simp: mono_def continuous_at_imp_continuous_at_within continuous_at_ereal) lemma ereal_SUP: "\SUP a\A. ereal (f a)\ \ \ \ ereal (SUP a\A. f a) = (SUP a\A. ereal (f a))" by (simp add: ereal_Sup image_comp) lemma ereal_Inf: assumes *: "\INF a\A. ereal a\ \ \" shows "ereal (Inf A) = (INF a\A. ereal a)" proof (rule continuous_at_Inf_mono) obtain r where r: "ereal r = (INF a\A. ereal a)" "A \ {}" using * by (force simp: top_ereal_def) then show "bdd_below A" "A \ {}" by (auto intro!: INF_lower bdd_belowI[of _ r] simp flip: ereal_less_eq) qed (auto simp: mono_def continuous_at_imp_continuous_at_within continuous_at_ereal) lemma ereal_Inf': assumes *: "bdd_below A" "A \ {}" shows "ereal (Inf A) = (INF a\A. ereal a)" proof (rule ereal_Inf) from * obtain l u where "x \ A \ l \ x" "u \ A" for x by (auto simp: bdd_below_def) then have "l \ (INF x\A. ereal x)" "(INF x\A. ereal x) \ u" by (auto intro!: INF_greatest INF_lower) then show "\INF a\A. ereal a\ \ \" by auto qed lemma ereal_INF: "\INF a\A. ereal (f a)\ \ \ \ ereal (INF a\A. f a) = (INF a\A. ereal (f a))" by (simp add: ereal_Inf image_comp) lemma ereal_Sup_uminus_image_eq: "Sup (uminus ` S::ereal set) = - Inf S" by (auto intro!: SUP_eqI simp: Ball_def[symmetric] ereal_uminus_le_reorder le_Inf_iff intro!: complete_lattice_class.Inf_lower2) lemma ereal_SUP_uminus_eq: fixes f :: "'a \ ereal" shows "(SUP x\S. uminus (f x)) = - (INF x\S. f x)" using ereal_Sup_uminus_image_eq [of "f ` S"] by (simp add: image_comp) lemma ereal_inj_on_uminus[intro, simp]: "inj_on uminus (A :: ereal set)" by (auto intro!: inj_onI) lemma ereal_Inf_uminus_image_eq: "Inf (uminus ` S::ereal set) = - Sup S" using ereal_Sup_uminus_image_eq[of "uminus ` S"] by simp lemma ereal_INF_uminus_eq: fixes f :: "'a \ ereal" shows "(INF x\S. - f x) = - (SUP x\S. f x)" using ereal_Inf_uminus_image_eq [of "f ` S"] by (simp add: image_comp) lemma ereal_SUP_uminus: fixes f :: "'a \ ereal" shows "(SUP i \ R. - f i) = - (INF i \ R. f i)" using ereal_Sup_uminus_image_eq[of "f`R"] by (simp add: image_image) lemma ereal_SUP_not_infty: fixes f :: "_ \ ereal" shows "A \ {} \ l \ -\ \ u \ \ \ \a\A. l \ f a \ f a \ u \ \Sup (f ` A)\ \ \" using SUP_upper2[of _ A l f] SUP_least[of A f u] by (cases "Sup (f ` A)") auto lemma ereal_INF_not_infty: fixes f :: "_ \ ereal" shows "A \ {} \ l \ -\ \ u \ \ \ \a\A. l \ f a \ f a \ u \ \Inf (f ` A)\ \ \" using INF_lower2[of _ A f u] INF_greatest[of A l f] by (cases "Inf (f ` A)") auto lemma ereal_image_uminus_shift: fixes X Y :: "ereal set" shows "uminus ` X = Y \ X = uminus ` Y" proof assume "uminus ` X = Y" then have "uminus ` uminus ` X = uminus ` Y" by (simp add: inj_image_eq_iff) then show "X = uminus ` Y" by (simp add: image_image) qed (simp add: image_image) lemma Sup_eq_MInfty: fixes S :: "ereal set" shows "Sup S = -\ \ S = {} \ S = {-\}" unfolding bot_ereal_def[symmetric] by auto lemma Inf_eq_PInfty: fixes S :: "ereal set" shows "Inf S = \ \ S = {} \ S = {\}" using Sup_eq_MInfty[of "uminus`S"] unfolding ereal_Sup_uminus_image_eq ereal_image_uminus_shift by simp lemma Inf_eq_MInfty: fixes S :: "ereal set" shows "-\ \ S \ Inf S = -\" unfolding bot_ereal_def[symmetric] by auto lemma Sup_eq_PInfty: fixes S :: "ereal set" shows "\ \ S \ Sup S = \" unfolding top_ereal_def[symmetric] by auto lemma not_MInfty_nonneg[simp]: "0 \ (x::ereal) \ x \ - \" by auto lemma Sup_ereal_close: fixes e :: ereal assumes "0 < e" and S: "\Sup S\ \ \" "S \ {}" shows "\x\S. Sup S - e < x" using assms by (cases e) (auto intro!: less_Sup_iff[THEN iffD1]) lemma Inf_ereal_close: fixes e :: ereal assumes "\Inf X\ \ \" and "0 < e" shows "\x\X. x < Inf X + e" proof (rule Inf_less_iff[THEN iffD1]) show "Inf X < Inf X + e" using assms by (cases e) auto qed lemma SUP_PInfty: "(\n::nat. \i\A. ereal (real n) \ f i) \ (SUP i\A. f i :: ereal) = \" unfolding top_ereal_def[symmetric] SUP_eq_top_iff by (metis MInfty_neq_PInfty(2) PInfty_neq_ereal(2) less_PInf_Ex_of_nat less_ereal.elims(2) less_le_trans) lemma SUP_nat_Infty: "(SUP i. ereal (real i)) = \" by (rule SUP_PInfty) auto lemma SUP_ereal_add_left: assumes "I \ {}" "c \ -\" shows "(SUP i\I. f i + c :: ereal) = (SUP i\I. f i) + c" proof (cases "(SUP i\I. f i) = - \") case True then have "\i. i \ I \ f i = -\" unfolding Sup_eq_MInfty by auto with True show ?thesis by (cases c) (auto simp: \I \ {}\) next case False then show ?thesis by (subst continuous_at_Sup_mono[where f="\x. x + c"]) (auto simp: continuous_at_imp_continuous_at_within continuous_at mono_def add_mono \I \ {}\ \c \ -\\ image_comp) qed lemma SUP_ereal_add_right: fixes c :: ereal shows "I \ {} \ c \ -\ \ (SUP i\I. c + f i) = c + (SUP i\I. f i)" using SUP_ereal_add_left[of I c f] by (simp add: add.commute) lemma SUP_ereal_minus_right: assumes "I \ {}" "c \ -\" shows "(SUP i\I. c - f i :: ereal) = c - (INF i\I. f i)" using SUP_ereal_add_right[OF assms, of "\i. - f i"] by (simp add: ereal_SUP_uminus minus_ereal_def) lemma SUP_ereal_minus_left: assumes "I \ {}" "c \ \" shows "(SUP i\I. f i - c:: ereal) = (SUP i\I. f i) - c" using SUP_ereal_add_left[OF \I \ {}\, of "-c" f] by (simp add: \c \ \\ minus_ereal_def) lemma INF_ereal_minus_right: assumes "I \ {}" and "\c\ \ \" shows "(INF i\I. c - f i) = c - (SUP i\I. f i::ereal)" proof - { fix b have "(-c) + b = - (c - b)" using \\c\ \ \\ by (cases c b rule: ereal2_cases) auto } note * = this show ?thesis using SUP_ereal_add_right[OF \I \ {}\, of "-c" f] \\c\ \ \\ by (auto simp add: * ereal_SUP_uminus_eq) qed lemma SUP_ereal_le_addI: fixes f :: "'i \ ereal" assumes "\i. f i + y \ z" and "y \ -\" shows "Sup (f ` UNIV) + y \ z" unfolding SUP_ereal_add_left[OF UNIV_not_empty \y \ -\\, symmetric] by (rule SUP_least assms)+ lemma SUP_combine: fixes f :: "'a::semilattice_sup \ 'a::semilattice_sup \ 'b::complete_lattice" assumes mono: "\a b c d. a \ b \ c \ d \ f a c \ f b d" shows "(SUP i\UNIV. SUP j\UNIV. f i j) = (SUP i. f i i)" proof (rule antisym) show "(SUP i j. f i j) \ (SUP i. f i i)" by (rule SUP_least SUP_upper2[where i="sup i j" for i j] UNIV_I mono sup_ge1 sup_ge2)+ show "(SUP i. f i i) \ (SUP i j. f i j)" by (rule SUP_least SUP_upper2 UNIV_I mono order_refl)+ qed lemma SUP_ereal_add: fixes f g :: "nat \ ereal" assumes inc: "incseq f" "incseq g" and pos: "\i. f i \ -\" "\i. g i \ -\" shows "(SUP i. f i + g i) = Sup (f ` UNIV) + Sup (g ` UNIV)" apply (subst SUP_ereal_add_left[symmetric, OF UNIV_not_empty]) apply (metis SUP_upper UNIV_I assms(4) ereal_infty_less_eq(2)) apply (subst (2) add.commute) apply (subst SUP_ereal_add_left[symmetric, OF UNIV_not_empty assms(3)]) apply (subst (2) add.commute) apply (rule SUP_combine[symmetric] add_mono inc[THEN monoD] | assumption)+ done lemma INF_eq_minf: "(INF i\I. f i::ereal) \ -\ \ (\b>-\. \i\I. b \ f i)" unfolding bot_ereal_def[symmetric] INF_eq_bot_iff by (auto simp: not_less) lemma INF_ereal_add_left: assumes "I \ {}" "c \ -\" "\x. x \ I \ 0 \ f x" shows "(INF i\I. f i + c :: ereal) = (INF i\I. f i) + c" proof - have "(INF i\I. f i) \ -\" unfolding INF_eq_minf using assms by (intro exI[of _ 0]) auto then show ?thesis by (subst continuous_at_Inf_mono[where f="\x. x + c"]) (auto simp: mono_def add_mono \I \ {}\ \c \ -\\ continuous_at_imp_continuous_at_within continuous_at image_comp) qed lemma INF_ereal_add_right: assumes "I \ {}" "c \ -\" "\x. x \ I \ 0 \ f x" shows "(INF i\I. c + f i :: ereal) = c + (INF i\I. f i)" using INF_ereal_add_left[OF assms] by (simp add: ac_simps) lemma INF_ereal_add_directed: fixes f g :: "'a \ ereal" assumes nonneg: "\i. i \ I \ 0 \ f i" "\i. i \ I \ 0 \ g i" assumes directed: "\i j. i \ I \ j \ I \ \k\I. f i + g j \ f k + g k" shows "(INF i\I. f i + g i) = (INF i\I. f i) + (INF i\I. g i)" proof cases assume "I = {}" then show ?thesis by (simp add: top_ereal_def) next assume "I \ {}" show ?thesis proof (rule antisym) show "(INF i\I. f i) + (INF i\I. g i) \ (INF i\I. f i + g i)" by (rule INF_greatest; intro add_mono INF_lower) next have "(INF i\I. f i + g i) \ (INF i\I. (INF j\I. f i + g j))" using directed by (intro INF_greatest) (blast intro: INF_lower2) also have "\ = (INF i\I. f i + (INF i\I. g i))" using nonneg \I \ {}\ by (auto simp: INF_ereal_add_right) also have "\ = (INF i\I. f i) + (INF i\I. g i)" using nonneg by (intro INF_ereal_add_left \I \ {}\) (auto simp: INF_eq_minf intro!: exI[of _ 0]) finally show "(INF i\I. f i + g i) \ (INF i\I. f i) + (INF i\I. g i)" . qed qed lemma INF_ereal_add: fixes f :: "nat \ ereal" assumes "decseq f" "decseq g" and fin: "\i. f i \ \" "\i. g i \ \" shows "(INF i. f i + g i) = Inf (f ` UNIV) + Inf (g ` UNIV)" proof - have INF_less: "(INF i. f i) < \" "(INF i. g i) < \" using assms unfolding INF_less_iff by auto { fix a b :: ereal assume "a \ \" "b \ \" then have "- ((- a) + (- b)) = a + b" by (cases a b rule: ereal2_cases) auto } note * = this have "(INF i. f i + g i) = (INF i. - ((- f i) + (- g i)))" by (simp add: fin *) also have "\ = Inf (f ` UNIV) + Inf (g ` UNIV)" unfolding ereal_INF_uminus_eq using assms INF_less by (subst SUP_ereal_add) (auto simp: ereal_SUP_uminus fin *) finally show ?thesis . qed lemma SUP_ereal_add_pos: fixes f g :: "nat \ ereal" assumes inc: "incseq f" "incseq g" and pos: "\i. 0 \ f i" "\i. 0 \ g i" shows "(SUP i. f i + g i) = Sup (f ` UNIV) + Sup (g ` UNIV)" proof (intro SUP_ereal_add inc) fix i show "f i \ -\" "g i \ -\" using pos[of i] by auto qed lemma SUP_ereal_sum: fixes f g :: "'a \ nat \ ereal" assumes "\n. n \ A \ incseq (f n)" and pos: "\n i. n \ A \ 0 \ f n i" shows "(SUP i. \n\A. f n i) = (\n\A. Sup ((f n) ` UNIV))" proof (cases "finite A") case True then show ?thesis using assms by induct (auto simp: incseq_sumI2 sum_nonneg SUP_ereal_add_pos) next case False then show ?thesis by simp qed lemma SUP_ereal_mult_left: fixes f :: "'a \ ereal" assumes "I \ {}" assumes f: "\i. i \ I \ 0 \ f i" and c: "0 \ c" shows "(SUP i\I. c * f i) = c * (SUP i\I. f i)" proof (cases "(SUP i \ I. f i) = 0") case True then have "\i. i \ I \ f i = 0" by (metis SUP_upper f antisym) with True show ?thesis by simp next case False then show ?thesis by (subst continuous_at_Sup_mono[where f="\x. c * x"]) (auto simp: mono_def continuous_at continuous_at_imp_continuous_at_within \I \ {}\ image_comp intro!: ereal_mult_left_mono c) qed lemma countable_approach: fixes x :: ereal assumes "x \ -\" shows "\f. incseq f \ (\i::nat. f i < x) \ (f \ x)" proof (cases x) case (real r) moreover have "(\n. r - inverse (real (Suc n))) \ r - 0" by (intro tendsto_intros LIMSEQ_inverse_real_of_nat) ultimately show ?thesis by (intro exI[of _ "\n. x - inverse (Suc n)"]) (auto simp: incseq_def) next case PInf with LIMSEQ_SUP[of "\n::nat. ereal (real n)"] show ?thesis by (intro exI[of _ "\n. ereal (real n)"]) (auto simp: incseq_def SUP_nat_Infty) qed (simp add: assms) lemma Sup_countable_SUP: assumes "A \ {}" shows "\f::nat \ ereal. incseq f \ range f \ A \ Sup A = (SUP i. f i)" proof cases assume "Sup A = -\" with \A \ {}\ have "A = {-\}" by (auto simp: Sup_eq_MInfty) then show ?thesis by (auto intro!: exI[of _ "\_. -\"] simp: bot_ereal_def) next assume "Sup A \ -\" then obtain l where "incseq l" and l: "l i < Sup A" and l_Sup: "l \ Sup A" for i :: nat by (auto dest: countable_approach) - have "\f. \n. (f n \ A \ l n \ f n) \ (f n \ f (Suc n))" + have "\f. \n. (f n \ A \ l n \ f n) \ (f n \ f (Suc n))" (is "\f. ?P f") proof (rule dependent_nat_choice) show "\x. x \ A \ l 0 \ x" using l[of 0] by (auto simp: less_Sup_iff) next fix x n assume "x \ A \ l n \ x" moreover from l[of "Suc n"] obtain y where "y \ A" "l (Suc n) < y" by (auto simp: less_Sup_iff) ultimately show "\y. (y \ A \ l (Suc n) \ y) \ x \ y" by (auto intro!: exI[of _ "max x y"] split: split_max) qed - then guess f .. note f = this + then obtain f where f: "?P f" .. then have "range f \ A" "incseq f" by (auto simp: incseq_Suc_iff) moreover have "(SUP i. f i) = Sup A" proof (rule tendsto_unique) show "f \ (SUP i. f i)" by (rule LIMSEQ_SUP \incseq f\)+ show "f \ Sup A" using l f by (intro tendsto_sandwich[OF _ _ l_Sup tendsto_const]) (auto simp: Sup_upper) qed simp ultimately show ?thesis by auto qed lemma Inf_countable_INF: assumes "A \ {}" shows "\f::nat \ ereal. decseq f \ range f \ A \ Inf A = (INF i. f i)" proof - obtain f where "incseq f" "range f \ uminus`A" "Sup (uminus`A) = (SUP i. f i)" using Sup_countable_SUP[of "uminus ` A"] \A \ {}\ by auto then show ?thesis by (intro exI[of _ "\x. - f x"]) (auto simp: ereal_Sup_uminus_image_eq ereal_INF_uminus_eq eq_commute[of "- _"]) qed lemma SUP_countable_SUP: "A \ {} \ \f::nat \ ereal. range f \ g`A \ Sup (g ` A) = Sup (f ` UNIV)" using Sup_countable_SUP [of "g`A"] by auto subsection "Relation to \<^typ>\enat\" definition "ereal_of_enat n = (case n of enat n \ ereal (real n) | \ \ \)" declare [[coercion "ereal_of_enat :: enat \ ereal"]] declare [[coercion "(\n. ereal (real n)) :: nat \ ereal"]] lemma ereal_of_enat_simps[simp]: "ereal_of_enat (enat n) = ereal n" "ereal_of_enat \ = \" by (simp_all add: ereal_of_enat_def) lemma ereal_of_enat_le_iff[simp]: "ereal_of_enat m \ ereal_of_enat n \ m \ n" by (cases m n rule: enat2_cases) auto lemma ereal_of_enat_less_iff[simp]: "ereal_of_enat m < ereal_of_enat n \ m < n" by (cases m n rule: enat2_cases) auto lemma numeral_le_ereal_of_enat_iff[simp]: "numeral m \ ereal_of_enat n \ numeral m \ n" by (cases n) (auto) lemma numeral_less_ereal_of_enat_iff[simp]: "numeral m < ereal_of_enat n \ numeral m < n" by (cases n) auto lemma ereal_of_enat_ge_zero_cancel_iff[simp]: "0 \ ereal_of_enat n \ 0 \ n" by (cases n) (auto simp flip: enat_0) lemma ereal_of_enat_gt_zero_cancel_iff[simp]: "0 < ereal_of_enat n \ 0 < n" by (cases n) (auto simp flip: enat_0) lemma ereal_of_enat_zero[simp]: "ereal_of_enat 0 = 0" by (auto simp flip: enat_0) lemma ereal_of_enat_inf[simp]: "ereal_of_enat n = \ \ n = \" by (cases n) auto lemma ereal_of_enat_add: "ereal_of_enat (m + n) = ereal_of_enat m + ereal_of_enat n" by (cases m n rule: enat2_cases) auto lemma ereal_of_enat_sub: assumes "n \ m" shows "ereal_of_enat (m - n) = ereal_of_enat m - ereal_of_enat n " using assms by (cases m n rule: enat2_cases) auto lemma ereal_of_enat_mult: "ereal_of_enat (m * n) = ereal_of_enat m * ereal_of_enat n" by (cases m n rule: enat2_cases) auto lemmas ereal_of_enat_pushin = ereal_of_enat_add ereal_of_enat_sub ereal_of_enat_mult lemmas ereal_of_enat_pushout = ereal_of_enat_pushin[symmetric] lemma ereal_of_enat_nonneg: "ereal_of_enat n \ 0" by(cases n) simp_all lemma ereal_of_enat_Sup: assumes "A \ {}" shows "ereal_of_enat (Sup A) = (SUP a \ A. ereal_of_enat a)" proof (intro antisym mono_Sup) show "ereal_of_enat (Sup A) \ (SUP a \ A. ereal_of_enat a)" proof cases assume "finite A" with \A \ {}\ obtain a where "a \ A" "ereal_of_enat (Sup A) = ereal_of_enat a" using Max_in[of A] by (auto simp: Sup_enat_def simp del: Max_in) then show ?thesis by (auto intro: SUP_upper) next assume "\ finite A" have [simp]: "(SUP a \ A. ereal_of_enat a) = top" unfolding SUP_eq_top_iff proof safe fix x :: ereal assume "x < top" then obtain n :: nat where "x < n" using less_PInf_Ex_of_nat top_ereal_def by auto obtain a where "a \ A - enat ` {.. n}" by (metis \\ finite A\ all_not_in_conv finite_Diff2 finite_atMost finite_imageI finite.emptyI) then have "a \ A" "ereal n \ ereal_of_enat a" by (auto simp: image_iff Ball_def) (metis enat_iless enat_ord_simps(1) ereal_of_enat_less_iff ereal_of_enat_simps(1) less_le not_less) with \x < n\ show "\i\A. x < ereal_of_enat i" by (auto intro!: bexI[of _ a]) qed show ?thesis by simp qed qed (simp add: mono_def) lemma ereal_of_enat_SUP: "A \ {} \ ereal_of_enat (SUP a\A. f a) = (SUP a \ A. ereal_of_enat (f a))" by (simp add: ereal_of_enat_Sup image_comp) subsection "Limits on \<^typ>\ereal\" lemma open_PInfty: "open A \ \ \ A \ (\x. {ereal x<..} \ A)" unfolding open_ereal_generated proof (induct rule: generate_topology.induct) case (Int A B) then obtain x z where "\ \ A \ {ereal x <..} \ A" "\ \ B \ {ereal z <..} \ B" by auto with Int show ?case by (intro exI[of _ "max x z"]) fastforce next case (Basis S) { fix x have "x \ \ \ \t. x \ ereal t" by (cases x) auto } moreover note Basis ultimately show ?case by (auto split: ereal.split) qed (fastforce simp add: vimage_Union)+ lemma open_MInfty: "open A \ -\ \ A \ (\x. {.. A)" unfolding open_ereal_generated proof (induct rule: generate_topology.induct) case (Int A B) then obtain x z where "-\ \ A \ {..< ereal x} \ A" "-\ \ B \ {..< ereal z} \ B" by auto with Int show ?case by (intro exI[of _ "min x z"]) fastforce next case (Basis S) { fix x have "x \ - \ \ \t. ereal t \ x" by (cases x) auto } moreover note Basis ultimately show ?case by (auto split: ereal.split) qed (fastforce simp add: vimage_Union)+ lemma open_ereal_vimage: "open S \ open (ereal -` S)" by (intro open_vimage continuous_intros) lemma open_ereal: "open S \ open (ereal ` S)" unfolding open_generated_order[where 'a=real] proof (induct rule: generate_topology.induct) case (Basis S) moreover have "\x. ereal ` {..< x} = { -\ <..< ereal x }" using ereal_less_ereal_Ex by auto moreover have "\x. ereal ` {x <..} = { ereal x <..< \ }" using less_ereal.elims(2) by fastforce ultimately show ?case by auto qed (auto simp add: image_Union image_Int) lemma open_image_real_of_ereal: fixes X::"ereal set" assumes "open X" assumes infty: "\ \ X" "-\ \ X" shows "open (real_of_ereal ` X)" proof - have "real_of_ereal ` X = ereal -` X" using infty ereal_real by (force simp: set_eq_iff) thus ?thesis by (auto intro!: open_ereal_vimage assms) qed lemma eventually_finite: fixes x :: ereal assumes "\x\ \ \" "(f \ x) F" shows "eventually (\x. \f x\ \ \) F" proof - have "(f \ ereal (real_of_ereal x)) F" using assms by (cases x) auto then have "eventually (\x. f x \ ereal ` UNIV) F" by (rule topological_tendstoD) (auto intro: open_ereal) also have "(\x. f x \ ereal ` UNIV) = (\x. \f x\ \ \)" by auto finally show ?thesis . qed lemma open_ereal_def: "open A \ open (ereal -` A) \ (\ \ A \ (\x. {ereal x <..} \ A)) \ (-\ \ A \ (\x. {.. A))" (is "open A \ ?rhs") proof assume "open A" then show ?rhs using open_PInfty open_MInfty open_ereal_vimage by auto next assume "?rhs" then obtain x y where A: "open (ereal -` A)" "\ \ A \ {ereal x<..} \ A" "-\ \ A \ {..< ereal y} \ A" by auto have *: "A = ereal ` (ereal -` A) \ (if \ \ A then {ereal x<..} else {}) \ (if -\ \ A then {..< ereal y} else {})" using A(2,3) by auto from open_ereal[OF A(1)] show "open A" by (subst *) (auto simp: open_Un) qed lemma open_PInfty2: assumes "open A" and "\ \ A" obtains x where "{ereal x<..} \ A" using open_PInfty[OF assms] by auto lemma open_MInfty2: assumes "open A" and "-\ \ A" obtains x where "{.. A" using open_MInfty[OF assms] by auto lemma ereal_openE: assumes "open A" obtains x y where "open (ereal -` A)" and "\ \ A \ {ereal x<..} \ A" and "-\ \ A \ {.. A" using assms open_ereal_def by auto lemmas open_ereal_lessThan = open_lessThan[where 'a=ereal] lemmas open_ereal_greaterThan = open_greaterThan[where 'a=ereal] lemmas ereal_open_greaterThanLessThan = open_greaterThanLessThan[where 'a=ereal] lemmas closed_ereal_atLeast = closed_atLeast[where 'a=ereal] lemmas closed_ereal_atMost = closed_atMost[where 'a=ereal] lemmas closed_ereal_atLeastAtMost = closed_atLeastAtMost[where 'a=ereal] lemmas closed_ereal_singleton = closed_singleton[where 'a=ereal] lemma ereal_open_cont_interval: fixes S :: "ereal set" assumes "open S" and "x \ S" and "\x\ \ \" obtains e where "e > 0" and "{x-e <..< x+e} \ S" proof - from \open S\ have "open (ereal -` S)" by (rule ereal_openE) then obtain e where "e > 0" and e: "dist y (real_of_ereal x) < e \ ereal y \ S" for y using assms unfolding open_dist by force show thesis proof (intro that subsetI) show "0 < ereal e" using \0 < e\ by auto fix y assume "y \ {x - ereal e<.. S" using e[of t] by auto qed qed lemma ereal_open_cont_interval2: fixes S :: "ereal set" assumes "open S" and "x \ S" and x: "\x\ \ \" obtains a b where "a < x" and "x < b" and "{a <..< b} \ S" proof - obtain e where "0 < e" "{x - e<.. S" using assms by (rule ereal_open_cont_interval) with that[of "x - e" "x + e"] ereal_between[OF x, of e] show thesis by auto qed subsubsection \Convergent sequences\ lemma lim_real_of_ereal[simp]: assumes lim: "(f \ ereal x) net" shows "((\x. real_of_ereal (f x)) \ x) net" proof (intro topological_tendstoI) fix S assume "open S" and "x \ S" then have S: "open S" "ereal x \ ereal ` S" by (simp_all add: inj_image_mem_iff) show "eventually (\x. real_of_ereal (f x) \ S) net" by (auto intro: eventually_mono [OF lim[THEN topological_tendstoD, OF open_ereal, OF S]]) qed lemma lim_ereal[simp]: "((\n. ereal (f n)) \ ereal x) net \ (f \ x) net" by (auto dest!: lim_real_of_ereal) lemma convergent_real_imp_convergent_ereal: assumes "convergent a" shows "convergent (\n. ereal (a n))" and "lim (\n. ereal (a n)) = ereal (lim a)" proof - from assms obtain L where L: "a \ L" unfolding convergent_def .. hence lim: "(\n. ereal (a n)) \ ereal L" using lim_ereal by auto thus "convergent (\n. ereal (a n))" unfolding convergent_def .. thus "lim (\n. ereal (a n)) = ereal (lim a)" using lim L limI by metis qed lemma tendsto_PInfty: "(f \ \) F \ (\r. eventually (\x. ereal r < f x) F)" proof - { fix l :: ereal assume "\r. eventually (\x. ereal r < f x) F" from this[THEN spec, of "real_of_ereal l"] have "l \ \ \ eventually (\x. l < f x) F" by (cases l) (auto elim: eventually_mono) } then show ?thesis by (auto simp: order_tendsto_iff) qed lemma tendsto_PInfty': "(f \ \) F = (\r>c. eventually (\x. ereal r < f x) F)" proof (subst tendsto_PInfty, intro iffI allI impI) assume A: "\r>c. eventually (\x. ereal r < f x) F" fix r :: real from A have A: "eventually (\x. ereal r < f x) F" if "r > c" for r using that by blast show "eventually (\x. ereal r < f x) F" proof (cases "r > c") case False hence B: "ereal r \ ereal (c + 1)" by simp have "c < c + 1" by simp from A[OF this] show "eventually (\x. ereal r < f x) F" by eventually_elim (rule le_less_trans[OF B]) qed (simp add: A) qed simp lemma tendsto_PInfty_eq_at_top: "((\z. ereal (f z)) \ \) F \ (LIM z F. f z :> at_top)" unfolding tendsto_PInfty filterlim_at_top_dense by simp lemma tendsto_MInfty: "(f \ -\) F \ (\r. eventually (\x. f x < ereal r) F)" unfolding tendsto_def proof safe fix S :: "ereal set" assume "open S" "-\ \ S" from open_MInfty[OF this] obtain B where "{.. S" .. moreover assume "\r::real. eventually (\z. f z < r) F" then have "eventually (\z. f z \ {..< B}) F" by auto ultimately show "eventually (\z. f z \ S) F" by (auto elim!: eventually_mono) next fix x assume "\S. open S \ -\ \ S \ eventually (\x. f x \ S) F" from this[rule_format, of "{..< ereal x}"] show "eventually (\y. f y < ereal x) F" by auto qed lemma tendsto_MInfty': "(f \ -\) F = (\rx. ereal r > f x) F)" proof (subst tendsto_MInfty, intro iffI allI impI) assume A: "\rx. ereal r > f x) F" fix r :: real from A have A: "eventually (\x. ereal r > f x) F" if "r < c" for r using that by blast show "eventually (\x. ereal r > f x) F" proof (cases "r < c") case False hence B: "ereal r \ ereal (c - 1)" by simp have "c > c - 1" by simp from A[OF this] show "eventually (\x. ereal r > f x) F" by eventually_elim (erule less_le_trans[OF _ B]) qed (simp add: A) qed simp lemma Lim_PInfty: "f \ \ \ (\B. \N. \n\N. f n \ ereal B)" unfolding tendsto_PInfty eventually_sequentially proof safe fix r assume "\r. \N. \n\N. ereal r \ f n" then obtain N where "\n\N. ereal (r + 1) \ f n" by blast moreover have "ereal r < ereal (r + 1)" by auto ultimately show "\N. \n\N. ereal r < f n" by (blast intro: less_le_trans) qed (blast intro: less_imp_le) lemma Lim_MInfty: "f \ -\ \ (\B. \N. \n\N. ereal B \ f n)" unfolding tendsto_MInfty eventually_sequentially proof safe fix r assume "\r. \N. \n\N. f n \ ereal r" then obtain N where "\n\N. f n \ ereal (r - 1)" by blast moreover have "ereal (r - 1) < ereal r" by auto ultimately show "\N. \n\N. f n < ereal r" by (blast intro: le_less_trans) qed (blast intro: less_imp_le) lemma Lim_bounded_PInfty: "f \ l \ (\n. f n \ ereal B) \ l \ \" using LIMSEQ_le_const2[of f l "ereal B"] by auto lemma Lim_bounded_MInfty: "f \ l \ (\n. ereal B \ f n) \ l \ -\" using LIMSEQ_le_const[of f l "ereal B"] by auto lemma tendsto_zero_erealI: assumes "\e. e > 0 \ eventually (\x. \f x\ < ereal e) F" shows "(f \ 0) F" proof (subst filterlim_cong[OF refl refl]) from assms[OF zero_less_one] show "eventually (\x. f x = ereal (real_of_ereal (f x))) F" by eventually_elim (auto simp: ereal_real) hence "eventually (\x. abs (real_of_ereal (f x)) < e) F" if "e > 0" for e using assms[OF that] by eventually_elim (simp add: real_less_ereal_iff that) hence "((\x. real_of_ereal (f x)) \ 0) F" unfolding tendsto_iff by (auto simp: tendsto_iff dist_real_def) thus "((\x. ereal (real_of_ereal (f x))) \ 0) F" by (simp add: zero_ereal_def) qed lemma Lim_bounded_PInfty2: "f \ l \ \n\N. f n \ ereal B \ l \ \" using LIMSEQ_le_const2[of f l "ereal B"] by fastforce lemma real_of_ereal_mult[simp]: fixes a b :: ereal shows "real_of_ereal (a * b) = real_of_ereal a * real_of_ereal b" by (cases rule: ereal2_cases[of a b]) auto lemma real_of_ereal_eq_0: fixes x :: ereal shows "real_of_ereal x = 0 \ x = \ \ x = -\ \ x = 0" by (cases x) auto lemma tendsto_ereal_realD: fixes f :: "'a \ ereal" assumes "x \ 0" and tendsto: "((\x. ereal (real_of_ereal (f x))) \ x) net" shows "(f \ x) net" proof (intro topological_tendstoI) fix S assume S: "open S" "x \ S" with \x \ 0\ have "open (S - {0})" "x \ S - {0}" by auto from tendsto[THEN topological_tendstoD, OF this] show "eventually (\x. f x \ S) net" by (rule eventually_rev_mp) (auto simp: ereal_real) qed lemma tendsto_ereal_realI: fixes f :: "'a \ ereal" assumes x: "\x\ \ \" and tendsto: "(f \ x) net" shows "((\x. ereal (real_of_ereal (f x))) \ x) net" proof (intro topological_tendstoI) fix S assume "open S" and "x \ S" with x have "open (S - {\, -\})" "x \ S - {\, -\}" by auto from tendsto[THEN topological_tendstoD, OF this] show "eventually (\x. ereal (real_of_ereal (f x)) \ S) net" by (elim eventually_mono) (auto simp: ereal_real) qed lemma ereal_mult_cancel_left: fixes a b c :: ereal shows "a * b = a * c \ (\a\ = \ \ 0 < b * c) \ a = 0 \ b = c" by (cases rule: ereal3_cases[of a b c]) (simp_all add: zero_less_mult_iff) lemma tendsto_add_ereal: fixes x y :: ereal assumes x: "\x\ \ \" and y: "\y\ \ \" assumes f: "(f \ x) F" and g: "(g \ y) F" shows "((\x. f x + g x) \ x + y) F" proof - from x obtain r where x': "x = ereal r" by (cases x) auto with f have "((\i. real_of_ereal (f i)) \ r) F" by simp moreover from y obtain p where y': "y = ereal p" by (cases y) auto with g have "((\i. real_of_ereal (g i)) \ p) F" by simp ultimately have "((\i. real_of_ereal (f i) + real_of_ereal (g i)) \ r + p) F" by (rule tendsto_add) moreover from eventually_finite[OF x f] eventually_finite[OF y g] have "eventually (\x. f x + g x = ereal (real_of_ereal (f x) + real_of_ereal (g x))) F" by eventually_elim auto ultimately show ?thesis by (simp add: x' y' cong: filterlim_cong) qed lemma tendsto_add_ereal_nonneg: fixes x y :: "ereal" assumes "x \ -\" "y \ -\" "(f \ x) F" "(g \ y) F" shows "((\x. f x + g x) \ x + y) F" proof cases assume "x = \ \ y = \" moreover { fix y :: ereal and f g :: "'a \ ereal" assume "y \ -\" "(f \ \) F" "(g \ y) F" then obtain y' where "-\ < y'" "y' < y" using dense[of "-\" y] by auto have "((\x. f x + g x) \ \) F" proof (rule tendsto_sandwich) have "\\<^sub>F x in F. y' < g x" using order_tendstoD(1)[OF \(g \ y) F\ \y' < y\] by auto then show "\\<^sub>F x in F. f x + y' \ f x + g x" by eventually_elim (auto intro!: add_mono) show "\\<^sub>F n in F. f n + g n \ \" "((\n. \) \ \) F" by auto show "((\x. f x + y') \ \) F" using tendsto_cadd_ereal[of y' \ f F] \(f \ \) F\ \-\ < y'\ by auto qed } note this[of y f g] this[of x g f] ultimately show ?thesis using assms by (auto simp: add_ac) next assume "\ (x = \ \ y = \)" with assms tendsto_add_ereal[of x y f F g] show ?thesis by auto qed lemma ereal_inj_affinity: fixes m t :: ereal assumes "\m\ \ \" and "m \ 0" and "\t\ \ \" shows "inj_on (\x. m * x + t) A" using assms by (cases rule: ereal2_cases[of m t]) (auto intro!: inj_onI simp: ereal_add_cancel_right ereal_mult_cancel_left) lemma ereal_PInfty_eq_plus[simp]: fixes a b :: ereal shows "\ = a + b \ a = \ \ b = \" by (cases rule: ereal2_cases[of a b]) auto lemma ereal_MInfty_eq_plus[simp]: fixes a b :: ereal shows "-\ = a + b \ (a = -\ \ b \ \) \ (b = -\ \ a \ \)" by (cases rule: ereal2_cases[of a b]) auto lemma ereal_less_divide_pos: fixes x y :: ereal shows "x > 0 \ x \ \ \ y < z / x \ x * y < z" by (cases rule: ereal3_cases[of x y z]) (auto simp: field_simps) lemma ereal_divide_less_pos: fixes x y z :: ereal shows "x > 0 \ x \ \ \ y / x < z \ y < x * z" by (cases rule: ereal3_cases[of x y z]) (auto simp: field_simps) lemma ereal_divide_eq: fixes a b c :: ereal shows "b \ 0 \ \b\ \ \ \ a / b = c \ a = b * c" by (cases rule: ereal3_cases[of a b c]) (simp_all add: field_simps) lemma ereal_inverse_not_MInfty[simp]: "inverse (a::ereal) \ -\" by (cases a) auto lemma ereal_mult_m1[simp]: "x * ereal (-1) = -x" by (cases x) auto lemma ereal_real': assumes "\x\ \ \" shows "ereal (real_of_ereal x) = x" using assms by auto lemma real_ereal_id: "real_of_ereal \ ereal = id" proof - { fix x have "(real_of_ereal \ ereal) x = id x" by auto } then show ?thesis using ext by blast qed lemma open_image_ereal: "open(UNIV-{ \ , (-\ :: ereal)})" by (metis range_ereal open_ereal open_UNIV) lemma ereal_le_distrib: fixes a b c :: ereal shows "c * (a + b) \ c * a + c * b" by (cases rule: ereal3_cases[of a b c]) (auto simp add: field_simps not_le mult_le_0_iff mult_less_0_iff) lemma ereal_pos_distrib: fixes a b c :: ereal assumes "0 \ c" and "c \ \" shows "c * (a + b) = c * a + c * b" using assms by (cases rule: ereal3_cases[of a b c]) (auto simp add: field_simps not_le mult_le_0_iff mult_less_0_iff) lemma ereal_LimI_finite: fixes x :: ereal assumes "\x\ \ \" and "\r. 0 < r \ \N. \n\N. u n < x + r \ x < u n + r" shows "u \ x" proof (rule topological_tendstoI, unfold eventually_sequentially) obtain rx where rx: "x = ereal rx" using assms by (cases x) auto fix S assume "open S" and "x \ S" then have "open (ereal -` S)" unfolding open_ereal_def by auto with \x \ S\ obtain r where "0 < r" and dist: "dist y rx < r \ ereal y \ S" for y unfolding open_dist rx by auto then obtain n where upper: "u N < x + ereal r" and lower: "x < u N + ereal r" if "n \ N" for N using assms(2)[of "ereal r"] by auto show "\N. \n\N. u n \ S" proof (safe intro!: exI[of _ n]) fix N assume "n \ N" from upper[OF this] lower[OF this] assms \0 < r\ have "u N \ {\,(-\)}" by auto then obtain ra where ra_def: "(u N) = ereal ra" by (cases "u N") auto then have "rx < ra + r" and "ra < rx + r" using rx assms \0 < r\ lower[OF \n \ N\] upper[OF \n \ N\] by auto then have "dist (real_of_ereal (u N)) rx < r" using rx ra_def by (auto simp: dist_real_def abs_diff_less_iff field_simps) from dist[OF this] show "u N \ S" using \u N \ {\, -\}\ by (auto simp: ereal_real split: if_split_asm) qed qed lemma tendsto_obtains_N: assumes "f \ f0" assumes "open S" and "f0 \ S" obtains N where "\n\N. f n \ S" using assms using tendsto_def using lim_explicit[of f f0] assms by auto lemma ereal_LimI_finite_iff: fixes x :: ereal assumes "\x\ \ \" shows "u \ x \ (\r. 0 < r \ (\N. \n\N. u n < x + r \ x < u n + r))" (is "?lhs \ ?rhs") proof assume lim: "u \ x" { fix r :: ereal assume "r > 0" then obtain N where "\n\N. u n \ {x - r <..< x + r}" apply (subst tendsto_obtains_N[of u x "{x - r <..< x + r}"]) using lim ereal_between[of x r] assms \r > 0\ apply auto done then have "\N. \n\N. u n < x + r \ x < u n + r" using ereal_minus_less[of r x] by (cases r) auto } then show ?rhs by auto next assume ?rhs then show "u \ x" using ereal_LimI_finite[of x] assms by auto qed lemma ereal_Limsup_uminus: fixes f :: "'a \ ereal" shows "Limsup net (\x. - (f x)) = - Liminf net f" unfolding Limsup_def Liminf_def ereal_SUP_uminus ereal_INF_uminus_eq .. lemma liminf_bounded_iff: fixes x :: "nat \ ereal" shows "C \ liminf x \ (\BN. \n\N. B < x n)" (is "?lhs \ ?rhs") unfolding le_Liminf_iff eventually_sequentially .. lemma Liminf_add_le: fixes f g :: "_ \ ereal" assumes F: "F \ bot" assumes ev: "eventually (\x. 0 \ f x) F" "eventually (\x. 0 \ g x) F" shows "Liminf F f + Liminf F g \ Liminf F (\x. f x + g x)" unfolding Liminf_def proof (subst SUP_ereal_add_left[symmetric]) let ?F = "{P. eventually P F}" let ?INF = "\P g. Inf (g ` (Collect P))" show "?F \ {}" by (auto intro: eventually_True) show "(SUP P\?F. ?INF P g) \ - \" unfolding bot_ereal_def[symmetric] SUP_bot_conv INF_eq_bot_iff by (auto intro!: exI[of _ 0] ev simp: bot_ereal_def) have "(SUP P\?F. ?INF P f + (SUP P\?F. ?INF P g)) \ (SUP P\?F. (SUP P'\?F. ?INF P f + ?INF P' g))" proof (safe intro!: SUP_mono bexI[of _ "\x. P x \ 0 \ f x" for P]) fix P let ?P' = "\x. P x \ 0 \ f x" assume "eventually P F" with ev show "eventually ?P' F" by eventually_elim auto have "?INF P f + (SUP P\?F. ?INF P g) \ ?INF ?P' f + (SUP P\?F. ?INF P g)" by (intro add_mono INF_mono) auto also have "\ = (SUP P'\?F. ?INF ?P' f + ?INF P' g)" proof (rule SUP_ereal_add_right[symmetric]) show "Inf (f ` {x. P x \ 0 \ f x}) \ - \" unfolding bot_ereal_def[symmetric] INF_eq_bot_iff by (auto intro!: exI[of _ 0] ev simp: bot_ereal_def) qed fact finally show "?INF P f + (SUP P\?F. ?INF P g) \ (SUP P'\?F. ?INF ?P' f + ?INF P' g)" . qed also have "\ \ (SUP P\?F. INF x\Collect P. f x + g x)" proof (safe intro!: SUP_least) fix P Q assume *: "eventually P F" "eventually Q F" show "?INF P f + ?INF Q g \ (SUP P\?F. INF x\Collect P. f x + g x)" proof (rule SUP_upper2) show "(\x. P x \ Q x) \ ?F" using * by (auto simp: eventually_conj) show "?INF P f + ?INF Q g \ (INF x\{x. P x \ Q x}. f x + g x)" by (intro INF_greatest add_mono) (auto intro: INF_lower) qed qed finally show "(SUP P\?F. ?INF P f + (SUP P\?F. ?INF P g)) \ (SUP P\?F. INF x\Collect P. f x + g x)" . qed lemma Sup_ereal_mult_right': assumes nonempty: "Y \ {}" and x: "x \ 0" shows "(SUP i\Y. f i) * ereal x = (SUP i\Y. f i * ereal x)" (is "?lhs = ?rhs") proof(cases "x = 0") case True thus ?thesis by(auto simp add: nonempty zero_ereal_def[symmetric]) next case False show ?thesis proof(rule antisym) show "?rhs \ ?lhs" by(rule SUP_least)(simp add: ereal_mult_right_mono SUP_upper x) next have "?lhs / ereal x = (SUP i\Y. f i) * (ereal x / ereal x)" by(simp only: ereal_times_divide_eq) also have "\ = (SUP i\Y. f i)" using False by simp also have "\ \ ?rhs / x" proof(rule SUP_least) fix i assume "i \ Y" have "f i = f i * (ereal x / ereal x)" using False by simp also have "\ = f i * x / x" by(simp only: ereal_times_divide_eq) also from \i \ Y\ have "f i * x \ ?rhs" by(rule SUP_upper) hence "f i * x / x \ ?rhs / x" using x False by simp finally show "f i \ ?rhs / x" . qed finally have "(?lhs / x) * x \ (?rhs / x) * x" by(rule ereal_mult_right_mono)(simp add: x) also have "\ = ?rhs" using False ereal_divide_eq mult.commute by force also have "(?lhs / x) * x = ?lhs" using False ereal_divide_eq mult.commute by force finally show "?lhs \ ?rhs" . qed qed lemma Sup_ereal_mult_left': "\ Y \ {}; x \ 0 \ \ ereal x * (SUP i\Y. f i) = (SUP i\Y. ereal x * f i)" by(subst (1 2) mult.commute)(rule Sup_ereal_mult_right') lemma sup_continuous_add[order_continuous_intros]: fixes f g :: "'a::complete_lattice \ ereal" assumes nn: "\x. 0 \ f x" "\x. 0 \ g x" and cont: "sup_continuous f" "sup_continuous g" shows "sup_continuous (\x. f x + g x)" unfolding sup_continuous_def proof safe fix M :: "nat \ 'a" assume "incseq M" then show "f (SUP i. M i) + g (SUP i. M i) = (SUP i. f (M i) + g (M i))" using SUP_ereal_add_pos[of "\i. f (M i)" "\i. g (M i)"] nn cont[THEN sup_continuous_mono] cont[THEN sup_continuousD] by (auto simp: mono_def) qed lemma sup_continuous_mult_right[order_continuous_intros]: "0 \ c \ c < \ \ sup_continuous f \ sup_continuous (\x. f x * c :: ereal)" by (cases c) (auto simp: sup_continuous_def fun_eq_iff Sup_ereal_mult_right') lemma sup_continuous_mult_left[order_continuous_intros]: "0 \ c \ c < \ \ sup_continuous f \ sup_continuous (\x. c * f x :: ereal)" using sup_continuous_mult_right[of c f] by (simp add: mult_ac) lemma sup_continuous_ereal_of_enat[order_continuous_intros]: assumes f: "sup_continuous f" shows "sup_continuous (\x. ereal_of_enat (f x))" by (rule sup_continuous_compose[OF _ f]) (auto simp: sup_continuous_def ereal_of_enat_SUP) subsubsection \Sums\ lemma sums_ereal_positive: fixes f :: "nat \ ereal" assumes "\i. 0 \ f i" shows "f sums (SUP n. \ii. \j=0.. ereal" assumes "\i. 0 \ f i" shows "summable f" using sums_ereal_positive[of f, OF assms] unfolding summable_def by auto lemma sums_ereal: "(\x. ereal (f x)) sums ereal x \ f sums x" unfolding sums_def by simp lemma suminf_ereal_eq_SUP: fixes f :: "nat \ ereal" assumes "\i. 0 \ f i" shows "(\x. f x) = (SUP n. \i ereal" assumes "\N. (\n x" and pos: "\n. 0 \ f n" shows "suminf f \ x" proof (rule Lim_bounded) have "summable f" using pos[THEN summable_ereal_pos] . then show "(\N. \n suminf f" by (auto dest!: summable_sums simp: sums_def atLeast0LessThan) show "\n\0. sum f {.. x" using assms by auto qed lemma suminf_bound_add: fixes f :: "nat \ ereal" assumes "\N. (\n x" and pos: "\n. 0 \ f n" and "y \ -\" shows "suminf f + y \ x" proof (cases y) case (real r) then have "\N. (\n x - y" using assms by (simp add: ereal_le_minus) then have "(\ n. f n) \ x - y" using pos by (rule suminf_bound) then show "(\ n. f n) + y \ x" using assms real by (simp add: ereal_le_minus) qed (insert assms, auto) lemma suminf_upper: fixes f :: "nat \ ereal" assumes "\n. 0 \ f n" shows "(\n (\n. f n)" unfolding suminf_ereal_eq_SUP [OF assms] by (auto intro: complete_lattice_class.SUP_upper) lemma suminf_0_le: fixes f :: "nat \ ereal" assumes "\n. 0 \ f n" shows "0 \ (\n. f n)" using suminf_upper[of f 0, OF assms] by simp lemma suminf_le_pos: fixes f g :: "nat \ ereal" assumes "\N. f N \ g N" and "\N. 0 \ f N" shows "suminf f \ suminf g" proof (safe intro!: suminf_bound) fix n { fix N have "0 \ g N" using assms(2,1)[of N] by auto } have "sum f {.. sum g {.. \ suminf g" using \\N. 0 \ g N\ by (rule suminf_upper) finally show "sum f {.. suminf g" . qed (rule assms(2)) lemma suminf_half_series_ereal: "(\n. (1/2 :: ereal) ^ Suc n) = 1" using sums_ereal[THEN iffD2, OF power_half_series, THEN sums_unique, symmetric] by (simp add: one_ereal_def) lemma suminf_add_ereal: fixes f g :: "nat \ ereal" assumes "\i. 0 \ f i" "\i. 0 \ g i" shows "(\i. f i + g i) = suminf f + suminf g" proof - have "(SUP n. \i ereal" assumes "\i. 0 \ f i" and "0 \ a" shows "(\i. a * f i) = a * suminf f" by (auto simp: sum_ereal_right_distrib[symmetric] assms ereal_zero_le_0_iff sum_nonneg suminf_ereal_eq_SUP intro!: SUP_ereal_mult_left) lemma suminf_PInfty: fixes f :: "nat \ ereal" assumes "\i. 0 \ f i" and "suminf f \ \" shows "f i \ \" proof - from suminf_upper[of f "Suc i", OF assms(1)] assms(2) have "(\i \" by auto then show ?thesis unfolding sum_Pinfty by simp qed lemma suminf_PInfty_fun: assumes "\i. 0 \ f i" and "suminf f \ \" shows "\f'. f = (\x. ereal (f' x))" proof - have "\i. \r. f i = ereal r" proof fix i show "\r. f i = ereal r" using suminf_PInfty[OF assms] assms(1)[of i] by (cases "f i") auto qed from choice[OF this] show ?thesis by auto qed lemma summable_ereal: assumes "\i. 0 \ f i" and "(\i. ereal (f i)) \ \" shows "summable f" proof - have "0 \ (\i. ereal (f i))" using assms by (intro suminf_0_le) auto with assms obtain r where r: "(\i. ereal (f i)) = ereal r" by (cases "\i. ereal (f i)") auto from summable_ereal_pos[of "\x. ereal (f x)"] have "summable (\x. ereal (f x))" using assms by auto from summable_sums[OF this] have "(\x. ereal (f x)) sums (\x. ereal (f x))" by auto then show "summable f" unfolding r sums_ereal summable_def .. qed lemma suminf_ereal: assumes "\i. 0 \ f i" and "(\i. ereal (f i)) \ \" shows "(\i. ereal (f i)) = ereal (suminf f)" proof (rule sums_unique[symmetric]) from summable_ereal[OF assms] show "(\x. ereal (f x)) sums (ereal (suminf f))" unfolding sums_ereal using assms by (intro summable_sums summable_ereal) qed lemma suminf_ereal_minus: fixes f g :: "nat \ ereal" assumes ord: "\i. g i \ f i" "\i. 0 \ g i" and fin: "suminf f \ \" "suminf g \ \" shows "(\i. f i - g i) = suminf f - suminf g" proof - { fix i have "0 \ f i" using ord[of i] by auto } moreover from suminf_PInfty_fun[OF \\i. 0 \ f i\ fin(1)] obtain f' where [simp]: "f = (\x. ereal (f' x))" .. from suminf_PInfty_fun[OF \\i. 0 \ g i\ fin(2)] obtain g' where [simp]: "g = (\x. ereal (g' x))" .. { fix i have "0 \ f i - g i" using ord[of i] by (auto simp: ereal_le_minus_iff) } moreover have "suminf (\i. f i - g i) \ suminf f" using assms by (auto intro!: suminf_le_pos simp: field_simps) then have "suminf (\i. f i - g i) \ \" using fin by auto ultimately show ?thesis using assms \\i. 0 \ f i\ apply simp apply (subst (1 2 3) suminf_ereal) apply (auto intro!: suminf_diff[symmetric] summable_ereal) done qed lemma suminf_ereal_PInf [simp]: "(\x. \::ereal) = \" proof - have "(\i) \ (\x. \::ereal)" by (rule suminf_upper) auto then show ?thesis by simp qed lemma summable_real_of_ereal: fixes f :: "nat \ ereal" assumes f: "\i. 0 \ f i" and fin: "(\i. f i) \ \" shows "summable (\i. real_of_ereal (f i))" proof (rule summable_def[THEN iffD2]) have "0 \ (\i. f i)" using assms by (auto intro: suminf_0_le) with fin obtain r where r: "ereal r = (\i. f i)" by (cases "(\i. f i)") auto { fix i have "f i \ \" using f by (intro suminf_PInfty[OF _ fin]) auto then have "\f i\ \ \" using f[of i] by auto } note fin = this have "(\i. ereal (real_of_ereal (f i))) sums (\i. ereal (real_of_ereal (f i)))" using f by (auto intro!: summable_ereal_pos simp: ereal_le_real_iff zero_ereal_def) also have "\ = ereal r" using fin r by (auto simp: ereal_real) finally show "\r. (\i. real_of_ereal (f i)) sums r" by (auto simp: sums_ereal) qed lemma suminf_SUP_eq: fixes f :: "nat \ nat \ ereal" assumes "\i. incseq (\n. f n i)" and "\n i. 0 \ f n i" shows "(\i. SUP n. f n i) = (SUP n. \i. f n i)" proof - have *: "\n. (\ii _ \ ereal" assumes nonneg: "\i a. a \ A \ 0 \ f i a" shows "(\i. \a\A. f i a) = (\a\A. \i. f i a)" proof (cases "finite A") case True then show ?thesis using nonneg by induct (simp_all add: suminf_add_ereal sum_nonneg) next case False then show ?thesis by simp qed lemma suminf_ereal_eq_0: fixes f :: "nat \ ereal" assumes nneg: "\i. 0 \ f i" shows "(\i. f i) = 0 \ (\i. f i = 0)" proof assume "(\i. f i) = 0" { fix i assume "f i \ 0" with nneg have "0 < f i" by (auto simp: less_le) also have "f i = (\j. if j = i then f i else 0)" by (subst suminf_finite[where N="{i}"]) auto also have "\ \ (\i. f i)" using nneg by (auto intro!: suminf_le_pos) finally have False using \(\i. f i) = 0\ by auto } then show "\i. f i = 0" by auto qed simp lemma suminf_ereal_offset_le: fixes f :: "nat \ ereal" assumes f: "\i. 0 \ f i" shows "(\i. f (i + k)) \ suminf f" proof - have "(\n. \i (\i. f (i + k))" using summable_sums[OF summable_ereal_pos] by (simp add: sums_def atLeast0LessThan f) moreover have "(\n. \i (\i. f i)" using summable_sums[OF summable_ereal_pos] by (simp add: sums_def atLeast0LessThan f) then have "(\n. \i (\i. f i)" by (rule LIMSEQ_ignore_initial_segment) ultimately show ?thesis proof (rule LIMSEQ_le, safe intro!: exI[of _ k]) fix n assume "k \ n" have "(\ii plus k) i)" by (simp add: ac_simps) also have "\ = (\i\(plus k) ` {.. \ sum f {..i sum f {.. (\i. ereal (f i)) = ereal x" by (metis sums_ereal sums_unique) lemma suminf_ereal': "summable f \ (\i. ereal (f i)) = ereal (\i. f i)" by (metis sums_ereal sums_unique summable_def) lemma suminf_ereal_finite: "summable f \ (\i. ereal (f i)) \ \" by (auto simp: summable_def simp flip: sums_ereal sums_unique) lemma suminf_ereal_finite_neg: assumes "summable f" shows "(\x. ereal (f x)) \ -\" proof- from assms obtain x where "f sums x" by blast hence "(\x. ereal (f x)) sums ereal x" by (simp add: sums_ereal) from sums_unique[OF this] have "(\x. ereal (f x)) = ereal x" .. thus "(\x. ereal (f x)) \ -\" by simp_all qed lemma SUP_ereal_add_directed: fixes f g :: "'a \ ereal" assumes nonneg: "\i. i \ I \ 0 \ f i" "\i. i \ I \ 0 \ g i" assumes directed: "\i j. i \ I \ j \ I \ \k\I. f i + g j \ f k + g k" shows "(SUP i\I. f i + g i) = (SUP i\I. f i) + (SUP i\I. g i)" proof cases assume "I = {}" then show ?thesis by (simp add: bot_ereal_def) next assume "I \ {}" show ?thesis proof (rule antisym) show "(SUP i\I. f i + g i) \ (SUP i\I. f i) + (SUP i\I. g i)" by (rule SUP_least; intro add_mono SUP_upper) next have "bot < (SUP i\I. g i)" using \I \ {}\ nonneg(2) by (auto simp: bot_ereal_def less_SUP_iff) then have "(SUP i\I. f i) + (SUP i\I. g i) = (SUP i\I. f i + (SUP i\I. g i))" by (intro SUP_ereal_add_left[symmetric] \I \ {}\) auto also have "\ = (SUP i\I. (SUP j\I. f i + g j))" using nonneg(1) \I \ {}\ by (simp add: SUP_ereal_add_right) also have "\ \ (SUP i\I. f i + g i)" using directed by (intro SUP_least) (blast intro: SUP_upper2) finally show "(SUP i\I. f i) + (SUP i\I. g i) \ (SUP i\I. f i + g i)" . qed qed lemma SUP_ereal_sum_directed: fixes f g :: "'a \ 'b \ ereal" assumes "I \ {}" assumes directed: "\N i j. N \ A \ i \ I \ j \ I \ \k\I. \n\N. f n i \ f n k \ f n j \ f n k" assumes nonneg: "\n i. i \ I \ n \ A \ 0 \ f n i" shows "(SUP i\I. \n\A. f n i) = (\n\A. SUP i\I. f n i)" proof - have "N \ A \ (SUP i\I. \n\N. f n i) = (\n\N. SUP i\I. f n i)" for N proof (induction N rule: infinite_finite_induct) case (insert n N) - moreover have "(SUP i\I. f n i + (\l\N. f l i)) = (SUP i\I. f n i) + (SUP i\I. \l\N. f l i)" + have "(SUP i\I. f n i + (\l\N. f l i)) = (SUP i\I. f n i) + (SUP i\I. \l\N. f l i)" proof (rule SUP_ereal_add_directed) fix i assume "i \ I" then show "0 \ f n i" "0 \ (\l\N. f l i)" using insert by (auto intro!: sum_nonneg nonneg) next fix i j assume "i \ I" "j \ I" - from directed[OF \insert n N \ A\ this] guess k .. - then show "\k\I. f n i + (\l\N. f l j) \ f n k + (\l\N. f l k)" - by (intro bexI[of _ k]) (auto intro!: add_mono sum_mono) + from directed[OF insert(4) this] + show "\k\I. f n i + (\l\N. f l j) \ f n k + (\l\N. f l k)" + by (auto intro!: add_mono sum_mono) qed - ultimately show ?case + with insert show ?case by simp qed (simp_all add: SUP_constant \I \ {}\) from this[of A] show ?thesis by simp qed lemma suminf_SUP_eq_directed: fixes f :: "_ \ nat \ ereal" assumes "I \ {}" assumes directed: "\N i j. i \ I \ j \ I \ finite N \ \k\I. \n\N. f i n \ f k n \ f j n \ f k n" assumes nonneg: "\n i. 0 \ f n i" shows "(\i. SUP n\I. f n i) = (SUP n\I. \i. f n i)" proof (subst (1 2) suminf_ereal_eq_SUP) show "\n i. 0 \ f n i" "\i. 0 \ (SUP n\I. f n i)" using \I \ {}\ nonneg by (auto intro: SUP_upper2) show "(SUP n. \iI. f n i) = (SUP n\I. SUP j. \i \r::rat. x < real_of_rat r \ real_of_rat r < y" proof (cases x y rule: ereal2_cases, simp_all) fix r q :: real assume "r < q" from Rats_dense_in_real[OF this] show "\x. r < real_of_rat x \ real_of_rat x < q" by (fastforce simp: Rats_def) next fix r :: real show "\x. r < real_of_rat x" "\x. real_of_rat x < r" using gt_ex[of r] lt_ex[of r] Rats_dense_in_real by (auto simp: Rats_def) qed lemma continuous_within_ereal[intro, simp]: "x \ A \ continuous (at x within A) ereal" using continuous_on_eq_continuous_within[of A ereal] by (auto intro: continuous_on_ereal continuous_on_id) lemma ereal_open_uminus: fixes S :: "ereal set" assumes "open S" shows "open (uminus ` S)" using \open S\[unfolded open_generated_order] proof induct have "range uminus = (UNIV :: ereal set)" by (auto simp: image_iff ereal_uminus_eq_reorder) then show "open (range uminus :: ereal set)" by simp qed (auto simp add: image_Union image_Int) lemma ereal_uminus_complement: fixes S :: "ereal set" shows "uminus ` (- S) = - uminus ` S" by (auto intro!: bij_image_Compl_eq surjI[of _ uminus] simp: bij_betw_def) lemma ereal_closed_uminus: fixes S :: "ereal set" assumes "closed S" shows "closed (uminus ` S)" using assms unfolding closed_def ereal_uminus_complement[symmetric] by (rule ereal_open_uminus) lemma ereal_open_affinity_pos: fixes S :: "ereal set" assumes "open S" and m: "m \ \" "0 < m" and t: "\t\ \ \" shows "open ((\x. m * x + t) ` S)" proof - have "continuous_on UNIV (\x. inverse m * (x + - t))" using m t by (intro continuous_at_imp_continuous_on ballI continuous_at[THEN iffD2]; force) then have "open ((\x. inverse m * (x + -t)) -` S)" using \open S\ open_vimage by blast also have "(\x. inverse m * (x + -t)) -` S = (\x. (x - t) / m) -` S" using m t by (auto simp: divide_ereal_def mult.commute minus_ereal_def simp flip: uminus_ereal.simps) also have "(\x. (x - t) / m) -` S = (\x. m * x + t) ` S" using m t by (simp add: set_eq_iff image_iff) (metis abs_ereal_less0 abs_ereal_uminus ereal_divide_eq ereal_eq_minus ereal_minus(7,8) ereal_minus_less_minus ereal_mult_eq_PInfty ereal_uminus_uminus ereal_zero_mult) finally show ?thesis . qed lemma ereal_open_affinity: fixes S :: "ereal set" assumes "open S" and m: "\m\ \ \" "m \ 0" and t: "\t\ \ \" shows "open ((\x. m * x + t) ` S)" proof cases assume "0 < m" then show ?thesis using ereal_open_affinity_pos[OF \open S\ _ _ t, of m] m by auto next assume "\ 0 < m" then have "0 < -m" using \m \ 0\ by (cases m) auto then have m: "-m \ \" "0 < -m" using \\m\ \ \\ by (auto simp: ereal_uminus_eq_reorder) from ereal_open_affinity_pos[OF ereal_open_uminus[OF \open S\] m t] show ?thesis unfolding image_image by simp qed lemma open_uminus_iff: fixes S :: "ereal set" shows "open (uminus ` S) \ open S" using ereal_open_uminus[of S] ereal_open_uminus[of "uminus ` S"] by auto lemma ereal_Liminf_uminus: fixes f :: "'a \ ereal" shows "Liminf net (\x. - (f x)) = - Limsup net f" using ereal_Limsup_uminus[of _ "(\x. - (f x))"] by auto lemma Liminf_PInfty: fixes f :: "'a \ ereal" assumes "\ trivial_limit net" shows "(f \ \) net \ Liminf net f = \" unfolding tendsto_iff_Liminf_eq_Limsup[OF assms] using Liminf_le_Limsup[OF assms, of f] by auto lemma Limsup_MInfty: fixes f :: "'a \ ereal" assumes "\ trivial_limit net" shows "(f \ -\) net \ Limsup net f = -\" unfolding tendsto_iff_Liminf_eq_Limsup[OF assms] using Liminf_le_Limsup[OF assms, of f] by auto lemma convergent_ereal: \ \RENAME\ fixes X :: "nat \ 'a :: {complete_linorder,linorder_topology}" shows "convergent X \ limsup X = liminf X" using tendsto_iff_Liminf_eq_Limsup[of sequentially] by (auto simp: convergent_def) lemma limsup_le_liminf_real: fixes X :: "nat \ real" and L :: real assumes 1: "limsup X \ L" and 2: "L \ liminf X" shows "X \ L" proof - from 1 2 have "limsup X \ liminf X" by auto hence 3: "limsup X = liminf X" by (simp add: Liminf_le_Limsup order_class.order.antisym) hence 4: "convergent (\n. ereal (X n))" by (subst convergent_ereal) hence "limsup X = lim (\n. ereal(X n))" by (rule convergent_limsup_cl) also from 1 2 3 have "limsup X = L" by auto finally have "lim (\n. ereal(X n)) = L" .. hence "(\n. ereal (X n)) \ L" using "4" convergent_LIMSEQ_iff by force thus ?thesis by simp qed lemma liminf_PInfty: fixes X :: "nat \ ereal" shows "X \ \ \ liminf X = \" by (metis Liminf_PInfty trivial_limit_sequentially) lemma limsup_MInfty: fixes X :: "nat \ ereal" shows "X \ -\ \ limsup X = -\" by (metis Limsup_MInfty trivial_limit_sequentially) lemma SUP_eq_LIMSEQ: assumes "mono f" shows "(SUP n. ereal (f n)) = ereal x \ f \ x" proof have inc: "incseq (\i. ereal (f i))" using \mono f\ unfolding mono_def incseq_def by auto { assume "f \ x" then have "(\i. ereal (f i)) \ ereal x" by auto from SUP_Lim[OF inc this] show "(SUP n. ereal (f n)) = ereal x" . next assume "(SUP n. ereal (f n)) = ereal x" with LIMSEQ_SUP[OF inc] show "f \ x" by auto } qed lemma liminf_ereal_cminus: fixes f :: "nat \ ereal" assumes "c \ -\" shows "liminf (\x. c - f x) = c - limsup f" proof (cases c) case PInf then show ?thesis by (simp add: Liminf_const) next case (real r) then show ?thesis by (simp add: liminf_SUP_INF limsup_INF_SUP INF_ereal_minus_right SUP_ereal_minus_right) qed (use \c \ -\\ in simp) subsubsection \Continuity\ lemma continuous_at_of_ereal: "\x0 :: ereal\ \ \ \ continuous (at x0) real_of_ereal" unfolding continuous_at by (rule lim_real_of_ereal) (simp add: ereal_real) lemma nhds_ereal: "nhds (ereal r) = filtermap ereal (nhds r)" by (simp add: filtermap_nhds_open_map open_ereal continuous_at_of_ereal) lemma at_ereal: "at (ereal r) = filtermap ereal (at r)" by (simp add: filter_eq_iff eventually_at_filter nhds_ereal eventually_filtermap) lemma at_left_ereal: "at_left (ereal r) = filtermap ereal (at_left r)" by (simp add: filter_eq_iff eventually_at_filter nhds_ereal eventually_filtermap) lemma at_right_ereal: "at_right (ereal r) = filtermap ereal (at_right r)" by (simp add: filter_eq_iff eventually_at_filter nhds_ereal eventually_filtermap) lemma shows at_left_PInf: "at_left \ = filtermap ereal at_top" and at_right_MInf: "at_right (-\) = filtermap ereal at_bot" unfolding filter_eq_iff eventually_filtermap eventually_at_top_dense eventually_at_bot_dense eventually_at_left[OF ereal_less(5)] eventually_at_right[OF ereal_less(6)] by (auto simp add: ereal_all_split ereal_ex_split) lemma ereal_tendsto_simps1: "((f \ real_of_ereal) \ y) (at_left (ereal x)) \ (f \ y) (at_left x)" "((f \ real_of_ereal) \ y) (at_right (ereal x)) \ (f \ y) (at_right x)" "((f \ real_of_ereal) \ y) (at_left (\::ereal)) \ (f \ y) at_top" "((f \ real_of_ereal) \ y) (at_right (-\::ereal)) \ (f \ y) at_bot" unfolding tendsto_compose_filtermap at_left_ereal at_right_ereal at_left_PInf at_right_MInf by (auto simp: filtermap_filtermap filtermap_ident) lemma ereal_tendsto_simps2: "((ereal \ f) \ ereal a) F \ (f \ a) F" "((ereal \ f) \ \) F \ (LIM x F. f x :> at_top)" "((ereal \ f) \ -\) F \ (LIM x F. f x :> at_bot)" unfolding tendsto_PInfty filterlim_at_top_dense tendsto_MInfty filterlim_at_bot_dense using lim_ereal by (simp_all add: comp_def) lemma inverse_infty_ereal_tendsto_0: "inverse \\\ (0::ereal)" proof - have **: "((\x. ereal (inverse x)) \ ereal 0) at_infinity" by (intro tendsto_intros tendsto_inverse_0) then have "((\x. if x = 0 then \ else ereal (inverse x)) \ 0) at_top" proof (rule filterlim_mono_eventually) show "nhds (ereal 0) \ nhds 0" by (simp add: zero_ereal_def) show "(at_top::real filter) \ at_infinity" by (simp add: at_top_le_at_infinity) qed auto then show ?thesis unfolding at_infty_ereal_eq_at_top tendsto_compose_filtermap[symmetric] comp_def by auto qed lemma inverse_ereal_tendsto_pos: fixes x :: ereal assumes "0 < x" shows "inverse \x\ inverse x" proof (cases x) case (real r) with \0 < x\ have **: "(\x. ereal (inverse x)) \r\ ereal (inverse r)" by (auto intro!: tendsto_inverse) from real \0 < x\ show ?thesis by (auto simp: at_ereal tendsto_compose_filtermap[symmetric] eventually_at_filter intro!: Lim_transform_eventually[OF **] t1_space_nhds) qed (insert \0 < x\, auto intro!: inverse_infty_ereal_tendsto_0) lemma inverse_ereal_tendsto_at_right_0: "(inverse \ \) (at_right (0::ereal))" unfolding tendsto_compose_filtermap[symmetric] at_right_ereal zero_ereal_def by (subst filterlim_cong[OF refl refl, where g="\x. ereal (inverse x)"]) (auto simp: eventually_at_filter tendsto_PInfty_eq_at_top filterlim_inverse_at_top_right) lemmas ereal_tendsto_simps = ereal_tendsto_simps1 ereal_tendsto_simps2 lemma continuous_at_iff_ereal: fixes f :: "'a::t2_space \ real" shows "continuous (at x0 within s) f \ continuous (at x0 within s) (ereal \ f)" unfolding continuous_within comp_def lim_ereal .. lemma continuous_on_iff_ereal: fixes f :: "'a::t2_space => real" assumes "open A" shows "continuous_on A f \ continuous_on A (ereal \ f)" unfolding continuous_on_def comp_def lim_ereal .. lemma continuous_on_real: "continuous_on (UNIV - {\, -\::ereal}) real_of_ereal" using continuous_at_of_ereal continuous_on_eq_continuous_at open_image_ereal by auto lemma continuous_on_iff_real: fixes f :: "'a::t2_space \ ereal" assumes "\x. x \ A \ \f x\ \ \" shows "continuous_on A f \ continuous_on A (real_of_ereal \ f)" proof assume L: "continuous_on A f" have "f ` A \ UNIV - {\, -\}" using assms by force then show "continuous_on A (real_of_ereal \ f)" by (meson L continuous_on_compose continuous_on_real continuous_on_subset) next assume R: "continuous_on A (real_of_ereal \ f)" then have "continuous_on A (ereal \ (real_of_ereal \ f))" by (meson continuous_at_iff_ereal continuous_on_eq_continuous_within) then show "continuous_on A f" using assms ereal_real' by auto qed lemma continuous_uminus_ereal [continuous_intros]: "continuous_on (A :: ereal set) uminus" unfolding continuous_on_def by (intro ballI tendsto_uminus_ereal[of "\x. x::ereal"]) simp lemma ereal_uminus_atMost [simp]: "uminus ` {..(a::ereal)} = {-a..}" proof (intro equalityI subsetI) fix x :: ereal assume "x \ {-a..}" hence "-(-x) \ uminus ` {..a}" by (intro imageI) (simp add: ereal_uminus_le_reorder) thus "x \ uminus ` {..a}" by simp qed auto lemma continuous_on_inverse_ereal [continuous_intros]: "continuous_on {0::ereal ..} inverse" unfolding continuous_on_def proof clarsimp fix x :: ereal assume "0 \ x" moreover have "at 0 within {0 ..} = at_right (0::ereal)" by (auto simp: filter_eq_iff eventually_at_filter le_less) moreover have "at x within {0 ..} = at x" if "0 < x" using that by (intro at_within_nhd[of _ "{0<..}"]) auto ultimately show "(inverse \ inverse x) (at x within {0..})" by (auto simp: le_less inverse_ereal_tendsto_at_right_0 inverse_ereal_tendsto_pos) qed lemma continuous_inverse_ereal_nonpos: "continuous_on ({..<0} :: ereal set) inverse" proof (subst continuous_on_cong[OF refl]) have "continuous_on {(0::ereal)<..} inverse" by (rule continuous_on_subset[OF continuous_on_inverse_ereal]) auto thus "continuous_on {..<(0::ereal)} (uminus \ inverse \ uminus)" by (intro continuous_intros) simp_all qed simp lemma tendsto_inverse_ereal: assumes "(f \ (c :: ereal)) F" assumes "eventually (\x. f x \ 0) F" shows "((\x. inverse (f x)) \ inverse c) F" by (cases "F = bot") (auto intro!: tendsto_lowerbound assms continuous_on_tendsto_compose[OF continuous_on_inverse_ereal]) subsubsection \liminf and limsup\ lemma Limsup_ereal_mult_right: assumes "F \ bot" "(c::real) \ 0" shows "Limsup F (\n. f n * ereal c) = Limsup F f * ereal c" proof (rule Limsup_compose_continuous_mono) from assms show "continuous_on UNIV (\a. a * ereal c)" using tendsto_cmult_ereal[of "ereal c" "\x. x" ] by (force simp: continuous_on_def mult_ac) qed (insert assms, auto simp: mono_def ereal_mult_right_mono) lemma Liminf_ereal_mult_right: assumes "F \ bot" "(c::real) \ 0" shows "Liminf F (\n. f n * ereal c) = Liminf F f * ereal c" proof (rule Liminf_compose_continuous_mono) from assms show "continuous_on UNIV (\a. a * ereal c)" using tendsto_cmult_ereal[of "ereal c" "\x. x" ] by (force simp: continuous_on_def mult_ac) qed (use assms in \auto simp: mono_def ereal_mult_right_mono\) lemma Liminf_ereal_mult_left: assumes "F \ bot" "(c::real) \ 0" shows "Liminf F (\n. ereal c * f n) = ereal c * Liminf F f" using Liminf_ereal_mult_right[OF assms] by (subst (1 2) mult.commute) lemma Limsup_ereal_mult_left: assumes "F \ bot" "(c::real) \ 0" shows "Limsup F (\n. ereal c * f n) = ereal c * Limsup F f" using Limsup_ereal_mult_right[OF assms] by (subst (1 2) mult.commute) lemma limsup_ereal_mult_right: "(c::real) \ 0 \ limsup (\n. f n * ereal c) = limsup f * ereal c" by (rule Limsup_ereal_mult_right) simp_all lemma limsup_ereal_mult_left: "(c::real) \ 0 \ limsup (\n. ereal c * f n) = ereal c * limsup f" by (subst (1 2) mult.commute, rule limsup_ereal_mult_right) simp_all lemma Limsup_add_ereal_right: "F \ bot \ abs c \ \ \ Limsup F (\n. g n + (c :: ereal)) = Limsup F g + c" by (rule Limsup_compose_continuous_mono) (auto simp: mono_def add_mono continuous_on_def) lemma Limsup_add_ereal_left: "F \ bot \ abs c \ \ \ Limsup F (\n. (c :: ereal) + g n) = c + Limsup F g" by (subst (1 2) add.commute) (rule Limsup_add_ereal_right) lemma Liminf_add_ereal_right: "F \ bot \ abs c \ \ \ Liminf F (\n. g n + (c :: ereal)) = Liminf F g + c" by (rule Liminf_compose_continuous_mono) (auto simp: mono_def add_mono continuous_on_def) lemma Liminf_add_ereal_left: "F \ bot \ abs c \ \ \ Liminf F (\n. (c :: ereal) + g n) = c + Liminf F g" by (subst (1 2) add.commute) (rule Liminf_add_ereal_right) lemma assumes "F \ bot" assumes nonneg: "eventually (\x. f x \ (0::ereal)) F" shows Liminf_inverse_ereal: "Liminf F (\x. inverse (f x)) = inverse (Limsup F f)" and Limsup_inverse_ereal: "Limsup F (\x. inverse (f x)) = inverse (Liminf F f)" proof - define inv where [abs_def]: "inv x = (if x \ 0 then \ else inverse x)" for x :: ereal have "continuous_on ({..0} \ {0..}) inv" unfolding inv_def by (intro continuous_on_If) (auto intro!: continuous_intros) also have "{..0} \ {0..} = (UNIV :: ereal set)" by auto finally have cont: "continuous_on UNIV inv" . have antimono: "antimono inv" unfolding inv_def antimono_def by (auto intro!: ereal_inverse_antimono) have "Liminf F (\x. inverse (f x)) = Liminf F (\x. inv (f x))" using nonneg by (auto intro!: Liminf_eq elim!: eventually_mono simp: inv_def) also have "... = inv (Limsup F f)" by (simp add: assms(1) Liminf_compose_continuous_antimono[OF cont antimono]) also from assms have "Limsup F f \ 0" by (intro le_Limsup) simp_all hence "inv (Limsup F f) = inverse (Limsup F f)" by (simp add: inv_def) finally show "Liminf F (\x. inverse (f x)) = inverse (Limsup F f)" . have "Limsup F (\x. inverse (f x)) = Limsup F (\x. inv (f x))" using nonneg by (auto intro!: Limsup_eq elim!: eventually_mono simp: inv_def) also have "... = inv (Liminf F f)" by (simp add: assms(1) Limsup_compose_continuous_antimono[OF cont antimono]) also from assms have "Liminf F f \ 0" by (intro Liminf_bounded) simp_all hence "inv (Liminf F f) = inverse (Liminf F f)" by (simp add: inv_def) finally show "Limsup F (\x. inverse (f x)) = inverse (Liminf F f)" . qed lemma ereal_diff_le_mono_left: "\ x \ z; 0 \ y \ \ x - y \ (z :: ereal)" by(cases x y z rule: ereal3_cases) simp_all lemma neg_0_less_iff_less_erea [simp]: "0 < - a \ (a :: ereal) < 0" by(cases a) simp_all lemma not_infty_ereal: "\x\ \ \ \ (\x'. x = ereal x')" by(cases x) simp_all lemma neq_PInf_trans: fixes x y :: ereal shows "\ y \ \; x \ y \ \ x \ \" by auto lemma mult_2_ereal: "ereal 2 * x = x + x" by(cases x) simp_all lemma ereal_diff_le_self: "0 \ y \ x - y \ (x :: ereal)" by(cases x y rule: ereal2_cases) simp_all lemma ereal_le_add_self: "0 \ y \ x \ x + (y :: ereal)" by(cases x y rule: ereal2_cases) simp_all lemma ereal_le_add_self2: "0 \ y \ x \ y + (x :: ereal)" by(cases x y rule: ereal2_cases) simp_all lemma ereal_le_add_mono1: "\ x \ y; 0 \ (z :: ereal) \ \ x \ y + z" using add_mono by fastforce lemma ereal_le_add_mono2: "\ x \ z; 0 \ (y :: ereal) \ \ x \ y + z" using add_mono by fastforce lemma ereal_diff_nonpos: fixes a b :: ereal shows "\ a \ b; a = \ \ b \ \; a = -\ \ b \ -\ \ \ a - b \ 0" by (cases rule: ereal2_cases[of a b]) auto lemma minus_ereal_0 [simp]: "x - ereal 0 = x" by(simp flip: zero_ereal_def) lemma ereal_diff_eq_0_iff: fixes a b :: ereal shows "(\a\ = \ \ \b\ \ \) \ a - b = 0 \ a = b" by(cases a b rule: ereal2_cases) simp_all lemma SUP_ereal_eq_0_iff_nonneg: fixes f :: "_ \ ereal" and A assumes nonneg: "\x\A. f x \ 0" and A:"A \ {}" shows "(SUP x\A. f x) = 0 \ (\x\A. f x = 0)" (is "?lhs \ ?rhs") proof(intro iffI ballI) fix x assume "?lhs" "x \ A" from \x \ A\ have "f x \ (SUP x\A. f x)" by(rule SUP_upper) with \?lhs\ show "f x = 0" using nonneg \x \ A\ by auto qed (simp add: A) lemma ereal_divide_le_posI: fixes x y z :: ereal shows "x > 0 \ z \ - \ \ z \ x * y \ z / x \ y" by (cases rule: ereal3_cases[of x y z])(auto simp: field_simps split: if_split_asm) lemma add_diff_eq_ereal: fixes x y z :: ereal shows "x + (y - z) = x + y - z" by(cases x y z rule: ereal3_cases) simp_all lemma ereal_diff_gr0: fixes a b :: ereal shows "a < b \ 0 < b - a" by (cases rule: ereal2_cases[of a b]) auto lemma ereal_minus_minus: fixes x y z :: ereal shows "(\y\ = \ \ \z\ \ \) \ x - (y - z) = x + z - y" by(cases x y z rule: ereal3_cases) simp_all lemma diff_add_eq_ereal: fixes a b c :: ereal shows "a - b + c = a + c - b" by(cases a b c rule: ereal3_cases) simp_all lemma diff_diff_commute_ereal: fixes x y z :: ereal shows "x - y - z = x - z - y" by(cases x y z rule: ereal3_cases) simp_all lemma ereal_diff_eq_MInfty_iff: fixes x y :: ereal shows "x - y = -\ \ x = -\ \ y \ -\ \ y = \ \ \x\ \ \" by(cases x y rule: ereal2_cases) simp_all lemma ereal_diff_add_inverse: fixes x y :: ereal shows "\x\ \ \ \ x + y - x = y" by(cases x y rule: ereal2_cases) simp_all lemma tendsto_diff_ereal: fixes x y :: ereal assumes x: "\x\ \ \" and y: "\y\ \ \" assumes f: "(f \ x) F" and g: "(g \ y) F" shows "((\x. f x - g x) \ x - y) F" proof - from x obtain r where x': "x = ereal r" by (cases x) auto with f have "((\i. real_of_ereal (f i)) \ r) F" by simp moreover from y obtain p where y': "y = ereal p" by (cases y) auto with g have "((\i. real_of_ereal (g i)) \ p) F" by simp ultimately have "((\i. real_of_ereal (f i) - real_of_ereal (g i)) \ r - p) F" by (rule tendsto_diff) moreover from eventually_finite[OF x f] eventually_finite[OF y g] have "eventually (\x. f x - g x = ereal (real_of_ereal (f x) - real_of_ereal (g x))) F" by eventually_elim auto ultimately show ?thesis by (simp add: x' y' cong: filterlim_cong) qed lemma continuous_on_diff_ereal: "continuous_on A f \ continuous_on A g \ (\x. x \ A \ \f x\ \ \) \ (\x. x \ A \ \g x\ \ \) \ continuous_on A (\z. f z - g z::ereal)" by (auto simp: tendsto_diff_ereal continuous_on_def) subsubsection \Tests for code generator\ text \A small list of simple arithmetic expressions.\ value "- \ :: ereal" value "\-\\ :: ereal" value "4 + 5 / 4 - ereal 2 :: ereal" value "ereal 3 < \" value "real_of_ereal (\::ereal) = 0" end