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 diff --git a/src/HOL/Library/Landau_Symbols.thy b/src/HOL/Library/Landau_Symbols.thy --- a/src/HOL/Library/Landau_Symbols.thy +++ b/src/HOL/Library/Landau_Symbols.thy @@ -1,2205 +1,2188 @@ (* File: Landau_Symbols_Definition.thy Author: Manuel Eberl - Landau symbols for reasoning about the asymptotic growth of functions. + Landau symbols for reasoning about the asymptotic growth of functions. *) section \Definition of Landau symbols\ theory Landau_Symbols -imports +imports Complex_Main begin lemma eventually_subst': "eventually (\x. f x = g x) F \ eventually (\x. P x (f x)) F = eventually (\x. P x (g x)) F" by (rule eventually_subst, erule eventually_rev_mp) simp subsection \Definition of Landau symbols\ text \ - Our Landau symbols are sign-oblivious, i.e. any function always has the same growth - as its absolute. This has the advantage of making some cancelling rules for sums nicer, but + Our Landau symbols are sign-oblivious, i.e. any function always has the same growth + as its absolute. This has the advantage of making some cancelling rules for sums nicer, but introduces some problems in other places. Nevertheless, we found this definition more convenient to work with. \ -definition bigo :: "'a filter \ ('a \ ('b :: real_normed_field)) \ ('a \ 'b) set" +definition bigo :: "'a filter \ ('a \ ('b :: real_normed_field)) \ ('a \ 'b) set" (\(1O[_]'(_'))\) - where "bigo F g = {f. (\c>0. eventually (\x. norm (f x) \ c * norm (g x)) F)}" + where "bigo F g = {f. (\c>0. eventually (\x. norm (f x) \ c * norm (g x)) F)}" -definition smallo :: "'a filter \ ('a \ ('b :: real_normed_field)) \ ('a \ 'b) set" +definition smallo :: "'a filter \ ('a \ ('b :: real_normed_field)) \ ('a \ 'b) set" (\(1o[_]'(_'))\) where "smallo F g = {f. (\c>0. eventually (\x. norm (f x) \ c * norm (g x)) F)}" -definition bigomega :: "'a filter \ ('a \ ('b :: real_normed_field)) \ ('a \ 'b) set" +definition bigomega :: "'a filter \ ('a \ ('b :: real_normed_field)) \ ('a \ 'b) set" (\(1\[_]'(_'))\) - where "bigomega F g = {f. (\c>0. eventually (\x. norm (f x) \ c * norm (g x)) F)}" + where "bigomega F g = {f. (\c>0. eventually (\x. norm (f x) \ c * norm (g x)) F)}" -definition smallomega :: "'a filter \ ('a \ ('b :: real_normed_field)) \ ('a \ 'b) set" +definition smallomega :: "'a filter \ ('a \ ('b :: real_normed_field)) \ ('a \ 'b) set" (\(1\[_]'(_'))\) where "smallomega F g = {f. (\c>0. eventually (\x. norm (f x) \ c * norm (g x)) F)}" -definition bigtheta :: "'a filter \ ('a \ ('b :: real_normed_field)) \ ('a \ 'b) set" +definition bigtheta :: "'a filter \ ('a \ ('b :: real_normed_field)) \ ('a \ 'b) set" (\(1\[_]'(_'))\) where "bigtheta F g = bigo F g \ bigomega F g" abbreviation bigo_at_top (\(2O'(_'))\) where - "O(g) \ bigo at_top g" + "O(g) \ bigo at_top g" abbreviation smallo_at_top (\(2o'(_'))\) where "o(g) \ smallo at_top g" abbreviation bigomega_at_top (\(2\'(_'))\) where "\(g) \ bigomega at_top g" abbreviation smallomega_at_top (\(2\'(_'))\) where "\(g) \ smallomega at_top g" abbreviation bigtheta_at_top (\(2\'(_'))\) where "\(g) \ bigtheta at_top g" - + text \The following is a set of properties that all Landau symbols satisfy.\ named_theorems landau_divide_simps locale landau_symbol = fixes L :: "'a filter \ ('a \ ('b :: real_normed_field)) \ ('a \ 'b) set" and L' :: "'c filter \ ('c \ ('b :: real_normed_field)) \ ('c \ 'b) set" and Lr :: "'a filter \ ('a \ real) \ ('a \ real) set" assumes bot': "L bot f = UNIV" assumes filter_mono': "F1 \ F2 \ L F2 f \ L F1 f" - assumes in_filtermap_iff: + assumes in_filtermap_iff: "f' \ L (filtermap h' F') g' \ (\x. f' (h' x)) \ L' F' (\x. g' (h' x))" - assumes filtercomap: + assumes filtercomap: "f' \ L F'' g' \ (\x. f' (h' x)) \ L' (filtercomap h' F'') (\x. g' (h' x))" assumes sup: "f \ L F1 g \ f \ L F2 g \ f \ L (sup F1 F2) g" assumes in_cong: "eventually (\x. f x = g x) F \ f \ L F (h) \ g \ L F (h)" assumes cong: "eventually (\x. f x = g x) F \ L F (f) = L F (g)" assumes cong_bigtheta: "f \ \[F](g) \ L F (f) = L F (g)" assumes in_cong_bigtheta: "f \ \[F](g) \ f \ L F (h) \ g \ L F (h)" assumes cmult [simp]: "c \ 0 \ L F (\x. c * f x) = L F (f)" assumes cmult_in_iff [simp]: "c \ 0 \ (\x. c * f x) \ L F (g) \ f \ L F (g)" assumes mult_left [simp]: "f \ L F (g) \ (\x. h x * f x) \ L F (\x. h x * g x)" - assumes inverse: "eventually (\x. f x \ 0) F \ eventually (\x. g x \ 0) F \ + assumes inverse: "eventually (\x. f x \ 0) F \ eventually (\x. g x \ 0) F \ f \ L F (g) \ (\x. inverse (g x)) \ L F (\x. inverse (f x))" assumes subsetI: "f \ L F (g) \ L F (f) \ L F (g)" assumes plus_subset1: "f \ o[F](g) \ L F (g) \ L F (\x. f x + g x)" assumes trans: "f \ L F (g) \ g \ L F (h) \ f \ L F (h)" assumes compose: "f \ L F (g) \ filterlim h' F G \ (\x. f (h' x)) \ L' G (\x. g (h' x))" assumes norm_iff [simp]: "(\x. norm (f x)) \ Lr F (\x. norm (g x)) \ f \ L F (g)" assumes abs [simp]: "Lr Fr (\x. \fr x\) = Lr Fr fr" assumes abs_in_iff [simp]: "(\x. \fr x\) \ Lr Fr gr \ fr \ Lr Fr gr" begin lemma bot [simp]: "f \ L bot g" by (simp add: bot') - + lemma filter_mono: "F1 \ F2 \ f \ L F2 g \ f \ L F1 g" using filter_mono'[of F1 F2] by blast -lemma cong_ex: +lemma cong_ex: "eventually (\x. f1 x = f2 x) F \ eventually (\x. g1 x = g2 x) F \ - f1 \ L F (g1) \ f2 \ L F (g2)" + f1 \ L F (g1) \ f2 \ L F (g2)" by (subst cong, assumption, subst in_cong, assumption, rule refl) -lemma cong_ex_bigtheta: - "f1 \ \[F](f2) \ g1 \ \[F](g2) \ f1 \ L F (g1) \ f2 \ L F (g2)" +lemma cong_ex_bigtheta: + "f1 \ \[F](f2) \ g1 \ \[F](g2) \ f1 \ L F (g1) \ f2 \ L F (g2)" by (subst cong_bigtheta, assumption, subst in_cong_bigtheta, assumption, rule refl) -lemma bigtheta_trans1: +lemma bigtheta_trans1: "f \ L F (g) \ g \ \[F](h) \ f \ L F (h)" by (subst cong_bigtheta[symmetric]) -lemma bigtheta_trans2: +lemma bigtheta_trans2: "f \ \[F](g) \ g \ L F (h) \ f \ L F (h)" by (subst in_cong_bigtheta) - + lemma cmult' [simp]: "c \ 0 \ L F (\x. f x * c) = L F (f)" by (subst mult.commute) (rule cmult) lemma cmult_in_iff' [simp]: "c \ 0 \ (\x. f x * c) \ L F (g) \ f \ L F (g)" by (subst mult.commute) (rule cmult_in_iff) lemma cdiv [simp]: "c \ 0 \ L F (\x. f x / c) = L F (f)" using cmult'[of "inverse c" F f] by (simp add: field_simps) lemma cdiv_in_iff' [simp]: "c \ 0 \ (\x. f x / c) \ L F (g) \ f \ L F (g)" using cmult_in_iff'[of "inverse c" f] by (simp add: field_simps) - + lemma uminus [simp]: "L F (\x. -g x) = L F (g)" using cmult[of "-1"] by simp lemma uminus_in_iff [simp]: "(\x. -f x) \ L F (g) \ f \ L F (g)" using cmult_in_iff[of "-1"] by simp lemma const: "c \ 0 \ L F (\_. c) = L F (\_. 1)" by (subst (2) cmult[symmetric]) simp_all (* Simplifier loops without the NO_MATCH *) lemma const' [simp]: "NO_MATCH 1 c \ c \ 0 \ L F (\_. c) = L F (\_. 1)" by (rule const) lemma const_in_iff: "c \ 0 \ (\_. c) \ L F (f) \ (\_. 1) \ L F (f)" using cmult_in_iff'[of c "\_. 1"] by simp lemma const_in_iff' [simp]: "NO_MATCH 1 c \ c \ 0 \ (\_. c) \ L F (f) \ (\_. 1) \ L F (f)" by (rule const_in_iff) lemma plus_subset2: "g \ o[F](f) \ L F (f) \ L F (\x. f x + g x)" by (subst add.commute) (rule plus_subset1) lemma mult_right [simp]: "f \ L F (g) \ (\x. f x * h x) \ L F (\x. g x * h x)" using mult_left by (simp add: mult.commute) lemma mult: "f1 \ L F (g1) \ f2 \ L F (g2) \ (\x. f1 x * f2 x) \ L F (\x. g1 x * g2 x)" by (rule trans, erule mult_left, erule mult_right) lemma inverse_cancel: assumes "eventually (\x. f x \ 0) F" assumes "eventually (\x. g x \ 0) F" shows "(\x. inverse (f x)) \ L F (\x. inverse (g x)) \ g \ L F (f)" proof assume "(\x. inverse (f x)) \ L F (\x. inverse (g x))" from inverse[OF _ _ this] assms show "g \ L F (f)" by simp qed (intro inverse assms) lemma divide_right: assumes "eventually (\x. h x \ 0) F" assumes "f \ L F (g)" shows "(\x. f x / h x) \ L F (\x. g x / h x)" by (subst (1 2) divide_inverse) (intro mult_right inverse assms) lemma divide_right_iff: assumes "eventually (\x. h x \ 0) F" shows "(\x. f x / h x) \ L F (\x. g x / h x) \ f \ L F (g)" proof assume "(\x. f x / h x) \ L F (\x. g x / h x)" from mult_right[OF this, of h] assms show "f \ L F (g)" by (subst (asm) cong_ex[of _ f F _ g]) (auto elim!: eventually_mono) qed (simp add: divide_right assms) lemma divide_left: assumes "eventually (\x. f x \ 0) F" assumes "eventually (\x. g x \ 0) F" assumes "g \ L F(f)" shows "(\x. h x / f x) \ L F (\x. h x / g x)" by (subst (1 2) divide_inverse) (intro mult_left inverse assms) lemma divide_left_iff: assumes "eventually (\x. f x \ 0) F" assumes "eventually (\x. g x \ 0) F" assumes "eventually (\x. h x \ 0) F" shows "(\x. h x / f x) \ L F (\x. h x / g x) \ g \ L F (f)" proof assume A: "(\x. h x / f x) \ L F (\x. h x / g x)" from assms have B: "eventually (\x. h x / f x / h x = inverse (f x)) F" by eventually_elim (simp add: divide_inverse) from assms have C: "eventually (\x. h x / g x / h x = inverse (g x)) F" by eventually_elim (simp add: divide_inverse) from divide_right[OF assms(3) A] assms show "g \ L F (f)" by (subst (asm) cong_ex[OF B C]) (simp add: inverse_cancel) qed (simp add: divide_left assms) lemma divide: assumes "eventually (\x. g1 x \ 0) F" assumes "eventually (\x. g2 x \ 0) F" assumes "f1 \ L F (f2)" "g2 \ L F (g1)" shows "(\x. f1 x / g1 x) \ L F (\x. f2 x / g2 x)" by (subst (1 2) divide_inverse) (intro mult inverse assms) - + lemma divide_eq1: assumes "eventually (\x. h x \ 0) F" shows "f \ L F (\x. g x / h x) \ (\x. f x * h x) \ L F (g)" proof- have "f \ L F (\x. g x / h x) \ (\x. f x * h x / h x) \ L F (\x. g x / h x)" using assms by (intro in_cong) (auto elim: eventually_mono) thus ?thesis by (simp only: divide_right_iff assms) qed lemma divide_eq2: assumes "eventually (\x. h x \ 0) F" shows "(\x. f x / h x) \ L F (\x. g x) \ f \ L F (\x. g x * h x)" proof- have "L F (\x. g x) = L F (\x. g x * h x / h x)" using assms by (intro cong) (auto elim: eventually_mono) thus ?thesis by (simp only: divide_right_iff assms) qed lemma inverse_eq1: assumes "eventually (\x. g x \ 0) F" shows "f \ L F (\x. inverse (g x)) \ (\x. f x * g x) \ L F (\_. 1)" using divide_eq1[of g F f "\_. 1"] by (simp add: divide_inverse assms) lemma inverse_eq2: assumes "eventually (\x. f x \ 0) F" shows "(\x. inverse (f x)) \ L F (g) \ (\x. 1) \ L F (\x. f x * g x)" using divide_eq2[of f F "\_. 1" g] by (simp add: divide_inverse assms mult_ac) lemma inverse_flip: assumes "eventually (\x. g x \ 0) F" assumes "eventually (\x. h x \ 0) F" assumes "(\x. inverse (g x)) \ L F (h)" shows "(\x. inverse (h x)) \ L F (g)" using assms by (simp add: divide_eq1 divide_eq2 inverse_eq_divide mult.commute) lemma lift_trans: assumes "f \ L F (g)" assumes "(\x. t x (g x)) \ L F (h)" assumes "\f g. f \ L F (g) \ (\x. t x (f x)) \ L F (\x. t x (g x))" shows "(\x. t x (f x)) \ L F (h)" by (rule trans[OF assms(3)[OF assms(1)] assms(2)]) lemma lift_trans': assumes "f \ L F (\x. t x (g x))" assumes "g \ L F (h)" assumes "\g h. g \ L F (h) \ (\x. t x (g x)) \ L F (\x. t x (h x))" shows "f \ L F (\x. t x (h x))" by (rule trans[OF assms(1) assms(3)[OF assms(2)]]) lemma lift_trans_bigtheta: assumes "f \ L F (g)" assumes "(\x. t x (g x)) \ \[F](h)" assumes "\f g. f \ L F (g) \ (\x. t x (f x)) \ L F (\x. t x (g x))" shows "(\x. t x (f x)) \ L F (h)" using cong_bigtheta[OF assms(2)] assms(3)[OF assms(1)] by simp lemma lift_trans_bigtheta': assumes "f \ L F (\x. t x (g x))" assumes "g \ \[F](h)" assumes "\g h. g \ \[F](h) \ (\x. t x (g x)) \ \[F](\x. t x (h x))" shows "f \ L F (\x. t x (h x))" using cong_bigtheta[OF assms(3)[OF assms(2)]] assms(1) by simp lemma (in landau_symbol) mult_in_1: assumes "f \ L F (\_. 1)" "g \ L F (\_. 1)" shows "(\x. f x * g x) \ L F (\_. 1)" using mult[OF assms] by simp lemma (in landau_symbol) of_real_cancel: "(\x. of_real (f x)) \ L F (\x. of_real (g x)) \ f \ Lr F g" by (subst (asm) norm_iff [symmetric], subst (asm) (1 2) norm_of_real) simp_all lemma (in landau_symbol) of_real_iff: "(\x. of_real (f x)) \ L F (\x. of_real (g x)) \ f \ Lr F g" by (subst norm_iff [symmetric], subst (1 2) norm_of_real) simp_all -lemmas [landau_divide_simps] = +lemmas [landau_divide_simps] = inverse_cancel divide_left_iff divide_eq1 divide_eq2 inverse_eq1 inverse_eq2 end text \ - The symbols $O$ and $o$ and $\Omega$ and $\omega$ are dual, so for many rules, replacing $O$ with + The symbols $O$ and $o$ and $\Omega$ and $\omega$ are dual, so for many rules, replacing $O$ with $\Omega$, $o$ with $\omega$, and $\leq$ with $\geq$ in a theorem yields another valid theorem. The following locale captures this fact. \ -locale landau_pair = +locale landau_pair = fixes L l :: "'a filter \ ('a \ ('b :: real_normed_field)) \ ('a \ 'b) set" fixes L' l' :: "'c filter \ ('c \ ('b :: real_normed_field)) \ ('c \ 'b) set" fixes Lr lr :: "'a filter \ ('a \ real) \ ('a \ real) set" and R :: "real \ real \ bool" assumes L_def: "L F g = {f. \c>0. eventually (\x. R (norm (f x)) (c * norm (g x))) F}" and l_def: "l F g = {f. \c>0. eventually (\x. R (norm (f x)) (c * norm (g x))) F}" and L'_def: "L' F' g' = {f. \c>0. eventually (\x. R (norm (f x)) (c * norm (g' x))) F'}" and l'_def: "l' F' g' = {f. \c>0. eventually (\x. R (norm (f x)) (c * norm (g' x))) F'}" and Lr_def: "Lr F'' g'' = {f. \c>0. eventually (\x. R (norm (f x)) (c * norm (g'' x))) F''}" and lr_def: "lr F'' g'' = {f. \c>0. eventually (\x. R (norm (f x)) (c * norm (g'' x))) F''}" and R: "R = (\) \ R = (\)" -interpretation landau_o: +interpretation landau_o: landau_pair bigo smallo bigo smallo bigo smallo "(\)" by unfold_locales (auto simp: bigo_def smallo_def intro!: ext) -interpretation landau_omega: +interpretation landau_omega: landau_pair bigomega smallomega bigomega smallomega bigomega smallomega "(\)" by unfold_locales (auto simp: bigomega_def smallomega_def intro!: ext) context landau_pair begin lemmas R_E = disjE [OF R, case_names le ge] lemma bigI: "c > 0 \ eventually (\x. R (norm (f x)) (c * norm (g x))) F \ f \ L F (g)" unfolding L_def by blast lemma bigE: assumes "f \ L F (g)" obtains c where "c > 0" "eventually (\x. R (norm (f x)) (c * (norm (g x)))) F" using assms unfolding L_def by blast lemma smallI: "(\c. c > 0 \ eventually (\x. R (norm (f x)) (c * (norm (g x)))) F) \ f \ l F (g)" unfolding l_def by blast lemma smallD: "f \ l F (g) \ c > 0 \ eventually (\x. R (norm (f x)) (c * (norm (g x)))) F" unfolding l_def by blast - + lemma bigE_nonneg_real: assumes "f \ Lr F (g)" "eventually (\x. f x \ 0) F" obtains c where "c > 0" "eventually (\x. R (f x) (c * \g x\)) F" proof- from assms(1) obtain c where c: "c > 0" "eventually (\x. R (norm (f x)) (c * norm (g x))) F" by (auto simp: Lr_def) from c(2) assms(2) have "eventually (\x. R (f x) (c * \g x\)) F" by eventually_elim simp from c(1) and this show ?thesis by (rule that) qed lemma smallD_nonneg_real: assumes "f \ lr F (g)" "eventually (\x. f x \ 0) F" "c > 0" shows "eventually (\x. R (f x) (c * \g x\)) F" using assms by (auto simp: lr_def dest!: spec[of _ c] elim: eventually_elim2) lemma small_imp_big: "f \ l F (g) \ f \ L F (g)" by (rule bigI[OF _ smallD, of 1]) simp_all - + lemma small_subset_big: "l F (g) \ L F (g)" using small_imp_big by blast lemma R_refl [simp]: "R x x" using R by auto lemma R_linear: "\R x y \ R y x" using R by auto lemma R_trans [trans]: "R a b \ R b c \ R a c" using R by auto lemma R_mult_left_mono: "R a b \ c \ 0 \ R (c*a) (c*b)" using R by (auto simp: mult_left_mono) lemma R_mult_right_mono: "R a b \ c \ 0 \ R (a*c) (b*c)" using R by (auto simp: mult_right_mono) lemma big_trans: assumes "f \ L F (g)" "g \ L F (h)" shows "f \ L F (h)" proof- - from assms(1) guess c by (elim bigE) note c = this - from assms(2) guess d by (elim bigE) note d = this - from c(2) d(2) have "eventually (\x. R (norm (f x)) (c * d * (norm (h x)))) F" + from assms obtain c d where *: "0 < c" "0 < d" + and **: "\\<^sub>F x in F. R (norm (f x)) (c * norm (g x))" + "\\<^sub>F x in F. R (norm (g x)) (d * norm (h x))" + by (elim bigE) + from ** have "eventually (\x. R (norm (f x)) (c * d * (norm (h x)))) F" proof eventually_elim fix x assume "R (norm (f x)) (c * (norm (g x)))" also assume "R (norm (g x)) (d * (norm (h x)))" - with c(1) have "R (c * (norm (g x))) (c * (d * (norm (h x))))" + with \0 < c\ have "R (c * (norm (g x))) (c * (d * (norm (h x))))" by (intro R_mult_left_mono) simp_all - finally show "R (norm (f x)) (c * d * (norm (h x)))" by (simp add: algebra_simps) + finally show "R (norm (f x)) (c * d * (norm (h x)))" + by (simp add: algebra_simps) qed - with c(1) d(1) show ?thesis by (intro bigI[of "c*d"]) simp_all + with * show ?thesis by (intro bigI[of "c*d"]) simp_all qed lemma big_small_trans: assumes "f \ L F (g)" "g \ l F (h)" shows "f \ l F (h)" proof (rule smallI) fix c :: real assume c: "c > 0" - from assms(1) guess d by (elim bigE) note d = this - note d(2) - moreover from c d assms(2) - have "eventually (\x. R (norm (g x)) (c * inverse d * norm (h x))) F" + from assms(1) obtain d where d: "d > 0" and *: "\\<^sub>F x in F. R (norm (f x)) (d * norm (g x))" + by (elim bigE) + from assms(2) c d have "eventually (\x. R (norm (g x)) (c * inverse d * norm (h x))) F" by (intro smallD) simp_all - ultimately show "eventually (\x. R (norm (f x)) (c * (norm (h x)))) F" - by eventually_elim (erule R_trans, insert R d(1), auto simp: field_simps) + with * show "eventually (\x. R (norm (f x)) (c * (norm (h x)))) F" + proof eventually_elim + case (elim x) + show ?case + by (use elim(1) in \rule R_trans\) (use elim(2) R d in \auto simp: field_simps\) + qed qed lemma small_big_trans: assumes "f \ l F (g)" "g \ L F (h)" shows "f \ l F (h)" proof (rule smallI) fix c :: real assume c: "c > 0" - from assms(2) guess d by (elim bigE) note d = this - note d(2) - moreover from c d assms(1) - have "eventually (\x. R (norm (f x)) (c * inverse d * norm (g x))) F" + from assms(2) obtain d where d: "d > 0" and *: "\\<^sub>F x in F. R (norm (g x)) (d * norm (h x))" + by (elim bigE) + from assms(1) c d have "eventually (\x. R (norm (f x)) (c * inverse d * norm (g x))) F" by (intro smallD) simp_all - ultimately show "eventually (\x. R (norm (f x)) (c * norm (h x))) F" - by eventually_elim (rotate_tac 2, erule R_trans, insert R c d(1), auto simp: field_simps) + with * show "eventually (\x. R (norm (f x)) (c * norm (h x))) F" + by eventually_elim (rotate_tac 2, erule R_trans, insert R c d, auto simp: field_simps) qed lemma small_trans: "f \ l F (g) \ g \ l F (h) \ f \ l F (h)" by (rule big_small_trans[OF small_imp_big]) lemma small_big_trans': "f \ l F (g) \ g \ L F (h) \ f \ L F (h)" by (rule small_imp_big[OF small_big_trans]) lemma big_small_trans': "f \ L F (g) \ g \ l F (h) \ f \ L F (h)" by (rule small_imp_big[OF big_small_trans]) lemma big_subsetI [intro]: "f \ L F (g) \ L F (f) \ L F (g)" by (intro subsetI) (drule (1) big_trans) lemma small_subsetI [intro]: "f \ L F (g) \ l F (f) \ l F (g)" by (intro subsetI) (drule (1) small_big_trans) lemma big_refl [simp]: "f \ L F (f)" by (rule bigI[of 1]) simp_all lemma small_refl_iff: "f \ l F (f) \ eventually (\x. f x = 0) F" proof (rule iffI[OF _ smallI]) assume f: "f \ l F f" have "(1/2::real) > 0" "(2::real) > 0" by simp_all from smallD[OF f this(1)] smallD[OF f this(2)] show "eventually (\x. f x = 0) F" by eventually_elim (insert R, auto) next fix c :: real assume "c > 0" "eventually (\x. f x = 0) F" from this(2) show "eventually (\x. R (norm (f x)) (c * norm (f x))) F" by eventually_elim simp_all qed lemma big_small_asymmetric: "f \ L F (g) \ g \ l F (f) \ eventually (\x. f x = 0) F" by (drule (1) big_small_trans) (simp add: small_refl_iff) lemma small_big_asymmetric: "f \ l F (g) \ g \ L F (f) \ eventually (\x. f x = 0) F" by (drule (1) small_big_trans) (simp add: small_refl_iff) lemma small_asymmetric: "f \ l F (g) \ g \ l F (f) \ eventually (\x. f x = 0) F" by (drule (1) small_trans) (simp add: small_refl_iff) lemma plus_aux: assumes "f \ o[F](g)" shows "g \ L F (\x. f x + g x)" proof (rule R_E) - assume [simp]: "R = (\)" + assume R: "R = (\)" have A: "1/2 > (0::real)" by simp - { - fix x assume "norm (f x) \ 1/2 * norm (g x)" - hence "1/2 * (norm (g x)) \ (norm (g x)) - (norm (f x))" by simp + have B: "1/2 * (norm (g x)) \ norm (f x + g x)" + if "norm (f x) \ 1/2 * norm (g x)" for x + proof - + from that have "1/2 * (norm (g x)) \ (norm (g x)) - (norm (f x))" + by simp also have "norm (g x) - norm (f x) \ norm (f x + g x)" by (subst add.commute) (rule norm_diff_ineq) - finally have "1/2 * (norm (g x)) \ norm (f x + g x)" by simp - } note B = this - + finally show ?thesis by simp + qed show "g \ L F (\x. f x + g x)" - apply (rule bigI[of "2"], simp) - using landau_o.smallD[OF assms A] apply eventually_elim - using B apply (simp add: algebra_simps) + apply (rule bigI[of "2"]) + apply simp + apply (use landau_o.smallD[OF assms A] in eventually_elim) + apply (use B in \simp add: R algebra_simps\) done next - assume [simp]: "R = (\x y. x \ y)" + assume R: "R = (\x y. x \ y)" show "g \ L F (\x. f x + g x)" proof (rule bigI[of "1/2"]) show "eventually (\x. R (norm (g x)) (1/2 * norm (f x + g x))) F" using landau_o.smallD[OF assms zero_less_one] proof eventually_elim case (elim x) - have "norm (f x + g x) \ norm (f x) + norm (g x)" by (rule norm_triangle_ineq) + have "norm (f x + g x) \ norm (f x) + norm (g x)" + by (rule norm_triangle_ineq) also note elim - finally show ?case by simp + finally show ?case by (simp add: R) qed qed simp_all qed end - lemma bigomega_iff_bigo: "g \ \[F](f) \ f \ O[F](g)" proof assume "f \ O[F](g)" - then guess c by (elim landau_o.bigE) - thus "g \ \[F](f)" by (intro landau_omega.bigI[of "inverse c"]) (simp_all add: field_simps) + then obtain c where "0 < c" "\\<^sub>F x in F. norm (f x) \ c * norm (g x)" + by (rule landau_o.bigE) + then show "g \ \[F](f)" + by (intro landau_omega.bigI[of "inverse c"]) (simp_all add: field_simps) next assume "g \ \[F](f)" - then guess c by (elim landau_omega.bigE) - thus "f \ O[F](g)" by (intro landau_o.bigI[of "inverse c"]) (simp_all add: field_simps) + then obtain c where "0 < c" "\\<^sub>F x in F. c * norm (f x) \ norm (g x)" + by (rule landau_omega.bigE) + then show "f \ O[F](g)" + by (intro landau_o.bigI[of "inverse c"]) (simp_all add: field_simps) qed lemma smallomega_iff_smallo: "g \ \[F](f) \ f \ o[F](g)" proof assume "f \ o[F](g)" from landau_o.smallD[OF this, of "inverse c" for c] show "g \ \[F](f)" by (intro landau_omega.smallI) (simp_all add: field_simps) next assume "g \ \[F](f)" from landau_omega.smallD[OF this, of "inverse c" for c] show "f \ o[F](g)" by (intro landau_o.smallI) (simp_all add: field_simps) qed context landau_pair begin lemma big_mono: "eventually (\x. R (norm (f x)) (norm (g x))) F \ f \ L F (g)" by (rule bigI[OF zero_less_one]) simp lemma big_mult: assumes "f1 \ L F (g1)" "f2 \ L F (g2)" shows "(\x. f1 x * f2 x) \ L F (\x. g1 x * g2 x)" proof- - from assms(1) guess c1 by (elim bigE) note c1 = this - from assms(2) guess c2 by (elim bigE) note c2 = this - - from c1(1) and c2(1) have "c1 * c2 > 0" by simp + from assms obtain c1 c2 where *: "c1 > 0" "c2 > 0" + and **: "\\<^sub>F x in F. R (norm (f1 x)) (c1 * norm (g1 x))" + "\\<^sub>F x in F. R (norm (f2 x)) (c2 * norm (g2 x))" + by (elim bigE) + from * have "c1 * c2 > 0" by simp moreover have "eventually (\x. R (norm (f1 x * f2 x)) (c1 * c2 * norm (g1 x * g2 x))) F" - using c1(2) c2(2) + using ** proof eventually_elim case (elim x) show ?case proof (cases rule: R_E) case le have "norm (f1 x) * norm (f2 x) \ (c1 * norm (g1 x)) * (c2 * norm (g2 x))" - using elim le c1(1) c2(1) by (intro mult_mono mult_nonneg_nonneg) auto + using elim le * by (intro mult_mono mult_nonneg_nonneg) auto with le show ?thesis by (simp add: le norm_mult mult_ac) next case ge have "(c1 * norm (g1 x)) * (c2 * norm (g2 x)) \ norm (f1 x) * norm (f2 x)" - using elim ge c1(1) c2(1) by (intro mult_mono mult_nonneg_nonneg) auto + using elim ge * by (intro mult_mono mult_nonneg_nonneg) auto with ge show ?thesis by (simp_all add: norm_mult mult_ac) qed qed ultimately show ?thesis by (rule bigI) qed lemma small_big_mult: assumes "f1 \ l F (g1)" "f2 \ L F (g2)" shows "(\x. f1 x * f2 x) \ l F (\x. g1 x * g2 x)" proof (rule smallI) fix c1 :: real assume c1: "c1 > 0" - from assms(2) guess c2 by (elim bigE) note c2 = this - with c1 assms(1) have "eventually (\x. R (norm (f1 x)) (c1 * inverse c2 * norm (g1 x))) F" + from assms(2) obtain c2 where c2: "c2 > 0" + and *: "\\<^sub>F x in F. R (norm (f2 x)) (c2 * norm (g2 x))" by (elim bigE) + from assms(1) c1 c2 have "eventually (\x. R (norm (f1 x)) (c1 * inverse c2 * norm (g1 x))) F" by (auto intro!: smallD) - thus "eventually (\x. R (norm (f1 x * f2 x)) (c1 * norm (g1 x * g2 x))) F" using c2(2) + with * show "eventually (\x. R (norm (f1 x * f2 x)) (c1 * norm (g1 x * g2 x))) F" proof eventually_elim case (elim x) show ?case proof (cases rule: R_E) case le have "norm (f1 x) * norm (f2 x) \ (c1 * inverse c2 * norm (g1 x)) * (c2 * norm (g2 x))" - using elim le c1(1) c2(1) by (intro mult_mono mult_nonneg_nonneg) auto - with le c2(1) show ?thesis by (simp add: le norm_mult field_simps) + using elim le c1 c2 by (intro mult_mono mult_nonneg_nonneg) auto + with le c2 show ?thesis by (simp add: le norm_mult field_simps) next case ge have "norm (f1 x) * norm (f2 x) \ (c1 * inverse c2 * norm (g1 x)) * (c2 * norm (g2 x))" - using elim ge c1(1) c2(1) by (intro mult_mono mult_nonneg_nonneg) auto - with ge c2(1) show ?thesis by (simp add: ge norm_mult field_simps) + using elim ge c1 c2 by (intro mult_mono mult_nonneg_nonneg) auto + with ge c2 show ?thesis by (simp add: ge norm_mult field_simps) qed qed qed -lemma big_small_mult: +lemma big_small_mult: "f1 \ L F (g1) \ f2 \ l F (g2) \ (\x. f1 x * f2 x) \ l F (\x. g1 x * g2 x)" by (subst (1 2) mult.commute) (rule small_big_mult) lemma small_mult: "f1 \ l F (g1) \ f2 \ l F (g2) \ (\x. f1 x * f2 x) \ l F (\x. g1 x * g2 x)" by (rule small_big_mult, assumption, rule small_imp_big) lemmas mult = big_mult small_big_mult big_small_mult small_mult sublocale big: landau_symbol L L' Lr proof have L: "L = bigo \ L = bigomega" by (rule R_E) (auto simp: bigo_def L_def bigomega_def fun_eq_iff) - { - fix c :: 'b and F and f :: "'a \ 'b" assume "c \ 0" - hence "(\x. c * f x) \ L F f" by (intro bigI[of "norm c"]) (simp_all add: norm_mult) - } note A = this - { - fix c :: 'b and F and f :: "'a \ 'b" assume "c \ 0" - from \c \ 0\ and A[of c f] and A[of "inverse c" "\x. c * f x"] - show "L F (\x. c * f x) = L F f" by (intro equalityI big_subsetI) (simp_all add: field_simps) - } - { - fix c :: 'b and F and f g :: "'a \ 'b" assume "c \ 0" + have A: "(\x. c * f x) \ L F f" if "c \ 0" for c :: 'b and F and f :: "'a \ 'b" + using that by (intro bigI[of "norm c"]) (simp_all add: norm_mult) + show "L F (\x. c * f x) = L F f" if "c \ 0" for c :: 'b and F and f :: "'a \ 'b" + using \c \ 0\ and A[of c f] and A[of "inverse c" "\x. c * f x"] + by (intro equalityI big_subsetI) (simp_all add: field_simps) + show "((\x. c * f x) \ L F g) = (f \ L F g)" if "c \ 0" + for c :: 'b and F and f g :: "'a \ 'b" + proof - from \c \ 0\ and A[of c f] and A[of "inverse c" "\x. c * f x"] - have "(\x. c * f x) \ L F f" "f \ L F (\x. c * f x)" by (simp_all add: field_simps) - thus "((\x. c * f x) \ L F g) = (f \ L F g)" by (intro iffI) (erule (1) big_trans)+ - } - { - fix f g :: "'a \ 'b" and F assume A: "f \ L F (g)" - assume B: "eventually (\x. f x \ 0) F" "eventually (\x. g x \ 0) F" - from A guess c by (elim bigE) note c = this - from c(2) B have "eventually (\x. R (norm (inverse (g x))) (c * norm (inverse (f x)))) F" - by eventually_elim (rule R_E, insert c(1), simp_all add: field_simps norm_inverse norm_divide) - with c(1) show "(\x. inverse (g x)) \ L F (\x. inverse (f x))" by (rule bigI) - } - { - fix f g :: "'a \ 'b" and F assume "f \ o[F](g)" - with plus_aux show "L F g \ L F (\x. f x + g x)" by (blast intro!: big_subsetI) - } - { - fix f g :: "'a \ 'b" and F assume A: "eventually (\x. f x = g x) F" - show "L F (f) = L F (g)" unfolding L_def - by (subst eventually_subst'[OF A]) (rule refl) - } - { - fix f g h :: "'a \ 'b" and F assume A: "eventually (\x. f x = g x) F" - show "f \ L F (h) \ g \ L F (h)" unfolding L_def mem_Collect_eq - by (subst (1) eventually_subst'[OF A]) (rule refl) - } - { - fix f g :: "'a \ 'b" and F assume "f \ L F g" thus "L F f \ L F g" by (rule big_subsetI) - } - { - fix f g :: "'a \ 'b" and F assume A: "f \ \[F](g)" - with A L show "L F (f) = L F (g)" unfolding bigtheta_def - by (intro equalityI big_subsetI) (auto simp: bigomega_iff_bigo) - fix h:: "'a \ 'b" - show "f \ L F (h) \ g \ L F (h)" by (rule disjE[OF L]) - (insert A, auto simp: bigtheta_def bigomega_iff_bigo intro: landau_o.big_trans) - } - { - fix f g h :: "'a \ 'b" and F assume "f \ L F g" - thus "(\x. h x * f x) \ L F (\x. h x * g x)" by (intro big_mult) simp - } - { - fix f g h :: "'a \ 'b" and F assume "f \ L F g" "g \ L F h" - thus "f \ L F (h)" by (rule big_trans) - } - { - fix f g :: "'a \ 'b" and h :: "'c \ 'a" and F G - assume "f \ L F g" and "filterlim h F G" - thus "(\x. f (h x)) \ L' G (\x. g (h x))" by (auto simp: L_def L'_def filterlim_iff) - } - { - fix f g :: "'a \ 'b" and F G :: "'a filter" - assume "f \ L F g" "f \ L G g" - from this [THEN bigE] guess c1 c2 . note c12 = this + have "(\x. c * f x) \ L F f" "f \ L F (\x. c * f x)" + by (simp_all add: field_simps) + then show ?thesis by (intro iffI) (erule (1) big_trans)+ + qed + show "(\x. inverse (g x)) \ L F (\x. inverse (f x))" + if *: "f \ L F (g)" and **: "eventually (\x. f x \ 0) F" "eventually (\x. g x \ 0) F" + for f g :: "'a \ 'b" and F + proof - + from * obtain c where c: "c > 0" and ***: "\\<^sub>F x in F. R (norm (f x)) (c * norm (g x))" + by (elim bigE) + from ** *** have "eventually (\x. R (norm (inverse (g x))) (c * norm (inverse (f x)))) F" + by eventually_elim (rule R_E, simp_all add: field_simps norm_inverse norm_divide c) + with c show ?thesis by (rule bigI) + qed + show "L F g \ L F (\x. f x + g x)" if "f \ o[F](g)" for f g :: "'a \ 'b" and F + using plus_aux that by (blast intro!: big_subsetI) + show "L F (f) = L F (g)" if "eventually (\x. f x = g x) F" for f g :: "'a \ 'b" and F + unfolding L_def by (subst eventually_subst'[OF that]) (rule refl) + show "f \ L F (h) \ g \ L F (h)" if "eventually (\x. f x = g x) F" + for f g h :: "'a \ 'b" and F + unfolding L_def mem_Collect_eq + by (subst (1) eventually_subst'[OF that]) (rule refl) + show "L F f \ L F g" if "f \ L F g" for f g :: "'a \ 'b" and F + using that by (rule big_subsetI) + show "L F (f) = L F (g)" if "f \ \[F](g)" for f g :: "'a \ 'b" and F + using L that unfolding bigtheta_def + by (intro equalityI big_subsetI) (auto simp: bigomega_iff_bigo) + show "f \ L F (h) \ g \ L F (h)" if "f \ \[F](g)" for f g h :: "'a \ 'b" and F + by (rule disjE[OF L]) + (use that in \auto simp: bigtheta_def bigomega_iff_bigo intro: landau_o.big_trans\) + show "(\x. h x * f x) \ L F (\x. h x * g x)" if "f \ L F g" for f g h :: "'a \ 'b" and F + using that by (intro big_mult) simp + show "f \ L F (h)" if "f \ L F g" "g \ L F h" for f g h :: "'a \ 'b" and F + using that by (rule big_trans) + show "(\x. f (h x)) \ L' G (\x. g (h x))" + if "f \ L F g" and "filterlim h F G" + for f g :: "'a \ 'b" and h :: "'c \ 'a" and F G + using that by (auto simp: L_def L'_def filterlim_iff) + show "f \ L (sup F G) g" if "f \ L F g" "f \ L G g" + for f g :: "'a \ 'b" and F G :: "'a filter" + proof - + from that [THEN bigE] obtain c1 c2 + where *: "c1 > 0" "c2 > 0" + and **: "\\<^sub>F x in F. R (norm (f x)) (c1 * norm (g x))" + "\\<^sub>F x in G. R (norm (f x)) (c2 * norm (g x))" . define c where "c = (if R c1 c2 then c2 else c1)" - from c12 have c: "R c1 c" "R c2 c" "c > 0" by (auto simp: c_def dest: R_linear) - with c12(2,4) have "eventually (\x. R (norm (f x)) (c * norm (g x))) F" - "eventually (\x. R (norm (f x)) (c * norm (g x))) G" + from * have c: "R c1 c" "R c2 c" "c > 0" + by (auto simp: c_def dest: R_linear) + with ** have "eventually (\x. R (norm (f x)) (c * norm (g x))) F" + "eventually (\x. R (norm (f x)) (c * norm (g x))) G" by (force elim: eventually_mono intro: R_trans[OF _ R_mult_right_mono])+ - with c show "f \ L (sup F G) g" by (auto simp: L_def eventually_sup) - } - { - fix f g :: "'a \ 'b" and h :: "'c \ 'a" and F G :: "'a filter" - assume "(f \ L F g)" - thus "((\x. f (h x)) \ L' (filtercomap h F) (\x. g (h x)))" - unfolding L_def L'_def by auto - } + with c show "f \ L (sup F G) g" + by (auto simp: L_def eventually_sup) + qed + show "((\x. f (h x)) \ L' (filtercomap h F) (\x. g (h x)))" if "(f \ L F g)" + for f g :: "'a \ 'b" and h :: "'c \ 'a" and F G :: "'a filter" + using that unfolding L_def L'_def by auto qed (auto simp: L_def Lr_def eventually_filtermap L'_def intro: filter_leD exI[of _ "1::real"]) sublocale small: landau_symbol l l' lr proof - { - fix c :: 'b and f :: "'a \ 'b" and F assume "c \ 0" - hence "(\x. c * f x) \ L F f" by (intro bigI[of "norm c"]) (simp_all add: norm_mult) - } note A = this - { - fix c :: 'b and f :: "'a \ 'b" and F assume "c \ 0" - from \c \ 0\ and A[of c f] and A[of "inverse c" "\x. c * f x"] - show "l F (\x. c * f x) = l F f" - by (intro equalityI small_subsetI) (simp_all add: field_simps) - } - { - fix c :: 'b and f g :: "'a \ 'b" and F assume "c \ 0" - from \c \ 0\ and A[of c f] and A[of "inverse c" "\x. c * f x"] - have "(\x. c * f x) \ L F f" "f \ L F (\x. c * f x)" by (simp_all add: field_simps) - thus "((\x. c * f x) \ l F g) = (f \ l F g)" by (intro iffI) (erule (1) big_small_trans)+ - } - { - fix f g :: "'a \ 'b" and F assume "f \ o[F](g)" - with plus_aux show "l F g \ l F (\x. f x + g x)" by (blast intro!: small_subsetI) - } - { - fix f g :: "'a \ 'b" and F assume A: "f \ l F (g)" - assume B: "eventually (\x. f x \ 0) F" "eventually (\x. g x \ 0) F" - show "(\x. inverse (g x)) \ l F (\x. inverse (f x))" - proof (rule smallI) - fix c :: real assume c: "c > 0" - from B smallD[OF A c] - show "eventually (\x. R (norm (inverse (g x))) (c * norm (inverse (f x)))) F" - by eventually_elim (rule R_E, simp_all add: field_simps norm_inverse norm_divide) - qed - } - { - fix f g :: "'a \ 'b" and F assume A: "eventually (\x. f x = g x) F" - show "l F (f) = l F (g)" unfolding l_def by (subst eventually_subst'[OF A]) (rule refl) - } - { - fix f g h :: "'a \ 'b" and F assume A: "eventually (\x. f x = g x) F" - show "f \ l F (h) \ g \ l F (h)" unfolding l_def mem_Collect_eq - by (subst (1) eventually_subst'[OF A]) (rule refl) - } - { - fix f g :: "'a \ 'b" and F assume "f \ l F g" - thus "l F f \ l F g" by (intro small_subsetI small_imp_big) - } - { - fix f g :: "'a \ 'b" and F assume A: "f \ \[F](g)" + have A: "(\x. c * f x) \ L F f" if "c \ 0" for c :: 'b and f :: "'a \ 'b" and F + using that by (intro bigI[of "norm c"]) (simp_all add: norm_mult) + show "l F (\x. c * f x) = l F f" if "c \ 0" for c :: 'b and f :: "'a \ 'b" and F + using that and A[of c f] and A[of "inverse c" "\x. c * f x"] + by (intro equalityI small_subsetI) (simp_all add: field_simps) + show "((\x. c * f x) \ l F g) = (f \ l F g)" if "c \ 0" for c :: 'b and f g :: "'a \ 'b" and F + proof - + from that and A[of c f] and A[of "inverse c" "\x. c * f x"] + have "(\x. c * f x) \ L F f" "f \ L F (\x. c * f x)" + by (simp_all add: field_simps) + then show ?thesis + by (intro iffI) (erule (1) big_small_trans)+ + qed + show "l F g \ l F (\x. f x + g x)" if "f \ o[F](g)" for f g :: "'a \ 'b" and F + using plus_aux that by (blast intro!: small_subsetI) + show "(\x. inverse (g x)) \ l F (\x. inverse (f x))" + if A: "f \ l F (g)" and B: "eventually (\x. f x \ 0) F" "eventually (\x. g x \ 0) F" + for f g :: "'a \ 'b" and F + proof (rule smallI) + fix c :: real assume c: "c > 0" + from B smallD[OF A c] + show "eventually (\x. R (norm (inverse (g x))) (c * norm (inverse (f x)))) F" + by eventually_elim (rule R_E, simp_all add: field_simps norm_inverse norm_divide) + qed + show "l F (f) = l F (g)" if "eventually (\x. f x = g x) F" for f g :: "'a \ 'b" and F + unfolding l_def by (subst eventually_subst'[OF that]) (rule refl) + show "f \ l F (h) \ g \ l F (h)" if "eventually (\x. f x = g x) F" for f g h :: "'a \ 'b" and F + unfolding l_def mem_Collect_eq by (subst (1) eventually_subst'[OF that]) (rule refl) + show "l F f \ l F g" if "f \ l F g" for f g :: "'a \ 'b" and F + using that by (intro small_subsetI small_imp_big) + show "l F (f) = l F (g)" if "f \ \[F](g)" for f g :: "'a \ 'b" and F + proof - have L: "L = bigo \ L = bigomega" by (rule R_E) (auto simp: bigo_def L_def bigomega_def fun_eq_iff) - with A show "l F (f) = l F (g)" unfolding bigtheta_def + with that show ?thesis unfolding bigtheta_def by (intro equalityI small_subsetI) (auto simp: bigomega_iff_bigo) + qed + show "f \ l F (h) \ g \ l F (h)" if "f \ \[F](g)" for f g h :: "'a \ 'b" and F + proof - have l: "l = smallo \ l = smallomega" by (rule R_E) (auto simp: smallo_def l_def smallomega_def fun_eq_iff) - fix h:: "'a \ 'b" - show "f \ l F (h) \ g \ l F (h)" by (rule disjE[OF l]) - (insert A, auto simp: bigtheta_def bigomega_iff_bigo smallomega_iff_smallo - intro: landau_o.big_small_trans landau_o.small_big_trans) - } - { - fix f g h :: "'a \ 'b" and F assume "f \ l F g" - thus "(\x. h x * f x) \ l F (\x. h x * g x)" by (intro big_small_mult) simp - } - { - fix f g h :: "'a \ 'b" and F assume "f \ l F g" "g \ l F h" - thus "f \ l F (h)" by (rule small_trans) - } - { - fix f g :: "'a \ 'b" and h :: "'c \ 'a" and F G - assume "f \ l F g" and "filterlim h F G" - thus "(\x. f (h x)) \ l' G (\x. g (h x))" - by (auto simp: l_def l'_def filterlim_iff) - } - { - fix f g :: "'a \ 'b" and h :: "'c \ 'a" and F G :: "'a filter" - assume "(f \ l F g)" - thus "((\x. f (h x)) \ l' (filtercomap h F) (\x. g (h x)))" - unfolding l_def l'_def by auto - } + show ?thesis + by (rule disjE[OF l]) + (use that in \auto simp: bigtheta_def bigomega_iff_bigo smallomega_iff_smallo + intro: landau_o.big_small_trans landau_o.small_big_trans\) + qed + show "(\x. h x * f x) \ l F (\x. h x * g x)" if "f \ l F g" for f g h :: "'a \ 'b" and F + using that by (intro big_small_mult) simp + show "f \ l F (h)" if "f \ l F g" "g \ l F h" for f g h :: "'a \ 'b" and F + using that by (rule small_trans) + show "(\x. f (h x)) \ l' G (\x. g (h x))" if "f \ l F g" and "filterlim h F G" + for f g :: "'a \ 'b" and h :: "'c \ 'a" and F G + using that by (auto simp: l_def l'_def filterlim_iff) + show "((\x. f (h x)) \ l' (filtercomap h F) (\x. g (h x)))" if "f \ l F g" + for f g :: "'a \ 'b" and h :: "'c \ 'a" and F G :: "'a filter" + using that unfolding l_def l'_def by auto qed (auto simp: l_def lr_def eventually_filtermap l'_def eventually_sup intro: filter_leD) text \These rules allow chaining of Landau symbol propositions in Isar with "also".\ lemma big_mult_1: "f \ L F (g) \ (\_. 1) \ L F (h) \ f \ L F (\x. g x * h x)" and big_mult_1': "(\_. 1) \ L F (g) \ f \ L F (h) \ f \ L F (\x. g x * h x)" and small_mult_1: "f \ l F (g) \ (\_. 1) \ L F (h) \ f \ l F (\x. g x * h x)" and small_mult_1': "(\_. 1) \ L F (g) \ f \ l F (h) \ f \ l F (\x. g x * h x)" and small_mult_1'': "f \ L F (g) \ (\_. 1) \ l F (h) \ f \ l F (\x. g x * h x)" and small_mult_1''': "(\_. 1) \ l F (g) \ f \ L F (h) \ f \ l F (\x. g x * h x)" by (drule (1) big.mult big_small_mult small_big_mult, simp)+ lemma big_1_mult: "f \ L F (g) \ h \ L F (\_. 1) \ (\x. f x * h x) \ L F (g)" and big_1_mult': "h \ L F (\_. 1) \ f \ L F (g) \ (\x. f x * h x) \ L F (g)" and small_1_mult: "f \ l F (g) \ h \ L F (\_. 1) \ (\x. f x * h x) \ l F (g)" and small_1_mult': "h \ L F (\_. 1) \ f \ l F (g) \ (\x. f x * h x) \ l F (g)" and small_1_mult'': "f \ L F (g) \ h \ l F (\_. 1) \ (\x. f x * h x) \ l F (g)" and small_1_mult''': "h \ l F (\_. 1) \ f \ L F (g) \ (\x. f x * h x) \ l F (g)" by (drule (1) big.mult big_small_mult small_big_mult, simp)+ -lemmas mult_1_trans = +lemmas mult_1_trans = big_mult_1 big_mult_1' small_mult_1 small_mult_1' small_mult_1'' small_mult_1''' big_1_mult big_1_mult' small_1_mult small_1_mult' small_1_mult'' small_1_mult''' lemma big_equal_iff_bigtheta: "L F (f) = L F (g) \ f \ \[F](g)" proof have L: "L = bigo \ L = bigomega" by (rule R_E) (auto simp: fun_eq_iff L_def bigo_def bigomega_def) fix f g :: "'a \ 'b" assume "L F (f) = L F (g)" with big_refl[of f F] big_refl[of g F] have "f \ L F (g)" "g \ L F (f)" by simp_all thus "f \ \[F](g)" using L unfolding bigtheta_def by (auto simp: bigomega_iff_bigo) qed (rule big.cong_bigtheta) lemma big_prod: assumes "\x. x \ A \ f x \ L F (g x)" shows "(\y. \x\A. f x y) \ L F (\y. \x\A. g x y)" using assms by (induction A rule: infinite_finite_induct) (auto intro!: big.mult) lemma big_prod_in_1: assumes "\x. x \ A \ f x \ L F (\_. 1)" shows "(\y. \x\A. f x y) \ L F (\_. 1)" using assms by (induction A rule: infinite_finite_induct) (auto intro!: big.mult_in_1) end context landau_symbol begin - + lemma plus_absorb1: assumes "f \ o[F](g)" shows "L F (\x. f x + g x) = L F (g)" proof (intro equalityI) from plus_subset1 and assms show "L F g \ L F (\x. f x + g x)" . from landau_o.small.plus_subset1[OF assms] and assms have "(\x. -f x) \ o[F](\x. f x + g x)" by (auto simp: landau_o.small.uminus_in_iff) from plus_subset1[OF this] show "L F (\x. f x + g x) \ L F (g)" by simp qed lemma plus_absorb2: "g \ o[F](f) \ L F (\x. f x + g x) = L F (f)" using plus_absorb1[of g F f] by (simp add: add.commute) lemma diff_absorb1: "f \ o[F](g) \ L F (\x. f x - g x) = L F (g)" by (simp only: diff_conv_add_uminus plus_absorb1 landau_o.small.uminus uminus) lemma diff_absorb2: "g \ o[F](f) \ L F (\x. f x - g x) = L F (f)" by (simp only: diff_conv_add_uminus plus_absorb2 landau_o.small.uminus_in_iff) lemmas absorb = plus_absorb1 plus_absorb2 diff_absorb1 diff_absorb2 end lemma bigthetaI [intro]: "f \ O[F](g) \ f \ \[F](g) \ f \ \[F](g)" unfolding bigtheta_def bigomega_def by blast -lemma bigthetaD1 [dest]: "f \ \[F](g) \ f \ O[F](g)" +lemma bigthetaD1 [dest]: "f \ \[F](g) \ f \ O[F](g)" and bigthetaD2 [dest]: "f \ \[F](g) \ f \ \[F](g)" unfolding bigtheta_def bigo_def bigomega_def by blast+ lemma bigtheta_refl [simp]: "f \ \[F](f)" unfolding bigtheta_def by simp lemma bigtheta_sym: "f \ \[F](g) \ g \ \[F](f)" unfolding bigtheta_def by (auto simp: bigomega_iff_bigo) lemmas landau_flip = bigomega_iff_bigo[symmetric] smallomega_iff_smallo[symmetric] bigomega_iff_bigo smallomega_iff_smallo bigtheta_sym interpretation landau_theta: landau_symbol bigtheta bigtheta bigtheta proof fix f g :: "'a \ 'b" and F assume "f \ o[F](g)" hence "O[F](g) \ O[F](\x. f x + g x)" "\[F](g) \ \[F](\x. f x + g x)" by (rule landau_o.big.plus_subset1 landau_omega.big.plus_subset1)+ thus "\[F](g) \ \[F](\x. f x + g x)" unfolding bigtheta_def by blast next - fix f g :: "'a \ 'b" and F + fix f g :: "'a \ 'b" and F assume "f \ \[F](g)" - thus A: "\[F](f) = \[F](g)" + thus A: "\[F](f) = \[F](g)" apply (subst (1 2) bigtheta_def) apply (subst landau_o.big.cong_bigtheta landau_omega.big.cong_bigtheta, assumption)+ apply (rule refl) done thus "\[F](f) \ \[F](g)" by simp fix h :: "'a \ 'b" show "f \ \[F](h) \ g \ \[F](h)" by (subst (1 2) bigtheta_sym) (simp add: A) next fix f g h :: "'a \ 'b" and F assume "f \ \[F](g)" "g \ \[F](h)" thus "f \ \[F](h)" unfolding bigtheta_def by (blast intro: landau_o.big.trans landau_omega.big.trans) next fix f :: "'a \ 'b" and F1 F2 :: "'a filter" assume "F1 \ F2" thus "\[F2](f) \ \[F1](f)" by (auto simp: bigtheta_def intro: landau_o.big.filter_mono landau_omega.big.filter_mono) -qed (auto simp: bigtheta_def landau_o.big.norm_iff - landau_o.big.cmult landau_omega.big.cmult - landau_o.big.cmult_in_iff landau_omega.big.cmult_in_iff +qed (auto simp: bigtheta_def landau_o.big.norm_iff + landau_o.big.cmult landau_omega.big.cmult + landau_o.big.cmult_in_iff landau_omega.big.cmult_in_iff landau_o.big.in_cong landau_omega.big.in_cong landau_o.big.mult landau_omega.big.mult - landau_o.big.inverse landau_omega.big.inverse + landau_o.big.inverse landau_omega.big.inverse landau_o.big.compose landau_omega.big.compose landau_o.big.bot' landau_omega.big.bot' landau_o.big.in_filtermap_iff landau_omega.big.in_filtermap_iff landau_o.big.sup landau_omega.big.sup landau_o.big.filtercomap landau_omega.big.filtercomap dest: landau_o.big.cong landau_omega.big.cong) -lemmas landau_symbols = +lemmas landau_symbols = landau_o.big.landau_symbol_axioms landau_o.small.landau_symbol_axioms - landau_omega.big.landau_symbol_axioms landau_omega.small.landau_symbol_axioms + landau_omega.big.landau_symbol_axioms landau_omega.small.landau_symbol_axioms landau_theta.landau_symbol_axioms lemma bigoI [intro]: assumes "eventually (\x. (norm (f x)) \ c * (norm (g x))) F" shows "f \ O[F](g)" proof (rule landau_o.bigI) show "max 1 c > 0" by simp - note assms - moreover have "\x. c * (norm (g x)) \ max 1 c * (norm (g x))" by (simp add: mult_right_mono) - ultimately show "eventually (\x. (norm (f x)) \ max 1 c * (norm (g x))) F" + have "c * (norm (g x)) \ max 1 c * (norm (g x))" for x + by (simp add: mult_right_mono) + with assms show "eventually (\x. (norm (f x)) \ max 1 c * (norm (g x))) F" by (auto elim!: eventually_mono dest: order.trans) qed lemma smallomegaD [dest]: assumes "f \ \[F](g)" shows "eventually (\x. (norm (f x)) \ c * (norm (g x))) F" proof (cases "c > 0") case False - show ?thesis + show ?thesis by (intro always_eventually allI, rule order.trans[of _ 0]) (insert False, auto intro!: mult_nonpos_nonneg) qed (blast dest: landau_omega.smallD[OF assms, of c]) - + lemma bigthetaI': assumes "c1 > 0" "c2 > 0" assumes "eventually (\x. c1 * (norm (g x)) \ (norm (f x)) \ (norm (f x)) \ c2 * (norm (g x))) F" shows "f \ \[F](g)" apply (rule bigthetaI) apply (rule landau_o.bigI[OF assms(2)]) using assms(3) apply (eventually_elim, simp) apply (rule landau_omega.bigI[OF assms(1)]) using assms(3) apply (eventually_elim, simp) done lemma bigthetaI_cong: "eventually (\x. f x = g x) F \ f \ \[F](g)" by (intro bigthetaI'[of 1 1]) (auto elim!: eventually_mono) -lemma (in landau_symbol) ev_eq_trans1: +lemma (in landau_symbol) ev_eq_trans1: "f \ L F (\x. g x (h x)) \ eventually (\x. h x = h' x) F \ f \ L F (\x. g x (h' x))" by (rule bigtheta_trans1[OF _ bigthetaI_cong]) (auto elim!: eventually_mono) -lemma (in landau_symbol) ev_eq_trans2: +lemma (in landau_symbol) ev_eq_trans2: "eventually (\x. f x = f' x) F \ (\x. g x (f' x)) \ L F (h) \ (\x. g x (f x)) \ L F (h)" by (rule bigtheta_trans2[OF bigthetaI_cong]) (auto elim!: eventually_mono) declare landau_o.smallI landau_omega.bigI landau_omega.smallI [intro] declare landau_o.bigE landau_omega.bigE [elim] declare landau_o.smallD -lemma (in landau_symbol) bigtheta_trans1': +lemma (in landau_symbol) bigtheta_trans1': "f \ L F (g) \ h \ \[F](g) \ f \ L F (h)" by (subst cong_bigtheta[symmetric]) (simp add: bigtheta_sym) -lemma (in landau_symbol) bigtheta_trans2': +lemma (in landau_symbol) bigtheta_trans2': "g \ \[F](f) \ g \ L F (h) \ f \ L F (h)" by (rule bigtheta_trans2, subst bigtheta_sym) lemma bigo_bigomega_trans: "f \ O[F](g) \ h \ \[F](g) \ f \ O[F](h)" and bigo_smallomega_trans: "f \ O[F](g) \ h \ \[F](g) \ f \ o[F](h)" and smallo_bigomega_trans: "f \ o[F](g) \ h \ \[F](g) \ f \ o[F](h)" and smallo_smallomega_trans: "f \ o[F](g) \ h \ \[F](g) \ f \ o[F](h)" and bigomega_bigo_trans: "f \ \[F](g) \ h \ O[F](g) \ f \ \[F](h)" and bigomega_smallo_trans: "f \ \[F](g) \ h \ o[F](g) \ f \ \[F](h)" and smallomega_bigo_trans: "f \ \[F](g) \ h \ O[F](g) \ f \ \[F](h)" and smallomega_smallo_trans: "f \ \[F](g) \ h \ o[F](g) \ f \ \[F](h)" by (unfold bigomega_iff_bigo smallomega_iff_smallo) - (erule (1) landau_o.big_trans landau_o.big_small_trans landau_o.small_big_trans + (erule (1) landau_o.big_trans landau_o.big_small_trans landau_o.small_big_trans landau_o.big_trans landau_o.small_trans)+ lemmas landau_trans_lift [trans] = landau_symbols[THEN landau_symbol.lift_trans] landau_symbols[THEN landau_symbol.lift_trans'] landau_symbols[THEN landau_symbol.lift_trans_bigtheta] landau_symbols[THEN landau_symbol.lift_trans_bigtheta'] lemmas landau_mult_1_trans [trans] = landau_o.mult_1_trans landau_omega.mult_1_trans -lemmas landau_trans [trans] = +lemmas landau_trans [trans] = landau_symbols[THEN landau_symbol.bigtheta_trans1] landau_symbols[THEN landau_symbol.bigtheta_trans2] landau_symbols[THEN landau_symbol.bigtheta_trans1'] landau_symbols[THEN landau_symbol.bigtheta_trans2'] landau_symbols[THEN landau_symbol.ev_eq_trans1] landau_symbols[THEN landau_symbol.ev_eq_trans2] landau_o.big_trans landau_o.small_trans landau_o.small_big_trans landau_o.big_small_trans - landau_omega.big_trans landau_omega.small_trans + landau_omega.big_trans landau_omega.small_trans landau_omega.small_big_trans landau_omega.big_small_trans - bigo_bigomega_trans bigo_smallomega_trans smallo_bigomega_trans smallo_smallomega_trans + bigo_bigomega_trans bigo_smallomega_trans smallo_bigomega_trans smallo_smallomega_trans bigomega_bigo_trans bigomega_smallo_trans smallomega_bigo_trans smallomega_smallo_trans -lemma bigtheta_inverse [simp]: +lemma bigtheta_inverse [simp]: shows "(\x. inverse (f x)) \ \[F](\x. inverse (g x)) \ f \ \[F](g)" -proof- - { - fix f g :: "'a \ 'b" and F assume A: "f \ \[F](g)" - then guess c1 c2 :: real unfolding bigtheta_def by (elim landau_o.bigE landau_omega.bigE IntE) - note c = this - from c(3) have "inverse c2 > 0" by simp - moreover from c(2,4) - have "eventually (\x. norm (inverse (f x)) \ inverse c2 * norm (inverse (g x))) F" +proof - + have "(\x. inverse (f x)) \ O[F](\x. inverse (g x))" + if A: "f \ \[F](g)" + for f g :: "'a \ 'b" and F + proof - + from A obtain c1 c2 :: real where *: "c1 > 0" "c2 > 0" + and **: "\\<^sub>F x in F. norm (f x) \ c1 * norm (g x)" + "\\<^sub>F x in F. c2 * norm (g x) \ norm (f x)" + unfolding bigtheta_def by (elim landau_o.bigE landau_omega.bigE IntE) + from \c2 > 0\ have c2: "inverse c2 > 0" by simp + from ** have "eventually (\x. norm (inverse (f x)) \ inverse c2 * norm (inverse (g x))) F" proof eventually_elim - fix x assume A: "(norm (f x)) \ c1 * (norm (g x))" "c2 * (norm (g x)) \ (norm (f x))" - from A c(1,3) have "f x = 0 \ g x = 0" by (auto simp: field_simps mult_le_0_iff) - with A c(1,3) show "norm (inverse (f x)) \ inverse c2 * norm (inverse (g x))" + fix x assume A: "norm (f x) \ c1 * norm (g x)" "c2 * norm (g x) \ norm (f x)" + from A * have "f x = 0 \ g x = 0" + by (auto simp: field_simps mult_le_0_iff) + with A * show "norm (inverse (f x)) \ inverse c2 * norm (inverse (g x))" by (force simp: field_simps norm_inverse norm_divide) qed - ultimately have "(\x. inverse (f x)) \ O[F](\x. inverse (g x))" by (rule landau_o.bigI) - } - thus ?thesis unfolding bigtheta_def + with c2 show ?thesis by (rule landau_o.bigI) + qed + then show ?thesis + unfolding bigtheta_def by (force simp: bigomega_iff_bigo bigtheta_sym) qed lemma bigtheta_divide: assumes "f1 \ \(f2)" "g1 \ \(g2)" shows "(\x. f1 x / g1 x) \ \(\x. f2 x / g2 x)" by (subst (1 2) divide_inverse, intro landau_theta.mult) (simp_all add: bigtheta_inverse assms) lemma eventually_nonzero_bigtheta: assumes "f \ \[F](g)" shows "eventually (\x. f x \ 0) F \ eventually (\x. g x \ 0) F" -proof- - { - fix f g :: "'a \ 'b" and F assume A: "f \ \[F](g)" and B: "eventually (\x. f x \ 0) F" - from A guess c1 c2 unfolding bigtheta_def by (elim landau_o.bigE landau_omega.bigE IntE) - from B this(2,4) have "eventually (\x. g x \ 0) F" by eventually_elim auto - } +proof - + have "eventually (\x. g x \ 0) F" + if A: "f \ \[F](g)" and B: "eventually (\x. f x \ 0) F" + for f g :: "'a \ 'b" + proof - + from A obtain c1 c2 where + "\\<^sub>F x in F. norm (f x) \ c1 * norm (g x)" + "\\<^sub>F x in F. c2 * norm (g x) \ norm (f x)" + unfolding bigtheta_def by (elim landau_o.bigE landau_omega.bigE IntE) + with B show ?thesis by eventually_elim auto + qed with assms show ?thesis by (force simp: bigtheta_sym) qed subsection \Landau symbols and limits\ lemma bigoI_tendsto_norm: fixes f g assumes "((\x. norm (f x / g x)) \ c) F" assumes "eventually (\x. g x \ 0) F" shows "f \ O[F](g)" proof (rule bigoI) - from assms have "eventually (\x. dist (norm (f x / g x)) c < 1) F" + from assms have "eventually (\x. dist (norm (f x / g x)) c < 1) F" using tendstoD by force thus "eventually (\x. (norm (f x)) \ (norm c + 1) * (norm (g x))) F" unfolding dist_real_def using assms(2) proof eventually_elim case (elim x) have "(norm (f x)) - norm c * (norm (g x)) \ norm ((norm (f x)) - c * (norm (g x)))" unfolding norm_mult [symmetric] using norm_triangle_ineq2[of "norm (f x)" "c * norm (g x)"] by (simp add: norm_mult abs_mult) also from elim have "\ = norm (norm (g x)) * norm (norm (f x / g x) - c)" unfolding norm_mult [symmetric] by (simp add: algebra_simps norm_divide) also from elim have "norm (norm (f x / g x) - c) \ 1" by simp - hence "norm (norm (g x)) * norm (norm (f x / g x) - c) \ norm (norm (g x)) * 1" + hence "norm (norm (g x)) * norm (norm (f x / g x) - c) \ norm (norm (g x)) * 1" by (rule mult_left_mono) simp_all finally show ?case by (simp add: algebra_simps) qed qed lemma bigoI_tendsto: assumes "((\x. f x / g x) \ c) F" assumes "eventually (\x. g x \ 0) F" shows "f \ O[F](g)" using assms by (rule bigoI_tendsto_norm[OF tendsto_norm]) lemma bigomegaI_tendsto_norm: assumes c_not_0: "(c::real) \ 0" assumes lim: "((\x. norm (f x / g x)) \ c) F" shows "f \ \[F](g)" proof (cases "F = bot") case False show ?thesis proof (rule landau_omega.bigI) from lim have "c \ 0" by (rule tendsto_lowerbound) (insert False, simp_all) with c_not_0 have "c > 0" by simp with c_not_0 show "c/2 > 0" by simp from lim have ev: "\\. \ > 0 \ eventually (\x. norm (norm (f x / g x) - c) < \) F" by (subst (asm) tendsto_iff) (simp add: dist_real_def) from ev[OF \c/2 > 0\] show "eventually (\x. (norm (f x)) \ c/2 * (norm (g x))) F" proof (eventually_elim) fix x assume B: "norm (norm (f x / g x) - c) < c / 2" from B have g: "g x \ 0" by auto from B have "-c/2 < -norm (norm (f x / g x) - c)" by simp also have "... \ norm (f x / g x) - c" by simp - finally show "(norm (f x)) \ c/2 * (norm (g x))" using g + finally show "(norm (f x)) \ c/2 * (norm (g x))" using g by (simp add: field_simps norm_mult norm_divide) qed qed qed simp lemma bigomegaI_tendsto: assumes c_not_0: "(c::real) \ 0" assumes lim: "((\x. f x / g x) \ c) F" shows "f \ \[F](g)" by (rule bigomegaI_tendsto_norm[OF _ tendsto_norm, of c]) (insert assms, simp_all) lemma smallomegaI_filterlim_at_top_norm: assumes lim: "filterlim (\x. norm (f x / g x)) at_top F" shows "f \ \[F](g)" proof (rule landau_omega.smallI) fix c :: real assume c_pos: "c > 0" from lim have ev: "eventually (\x. norm (f x / g x) \ c) F" by (subst (asm) filterlim_at_top) simp thus "eventually (\x. (norm (f x)) \ c * (norm (g x))) F" proof eventually_elim fix x assume A: "norm (f x / g x) \ c" from A c_pos have "g x \ 0" by auto with A show "(norm (f x)) \ c * (norm (g x))" by (simp add: field_simps norm_divide) qed qed lemma smallomegaI_filterlim_at_infinity: assumes lim: "filterlim (\x. f x / g x) at_infinity F" shows "f \ \[F](g)" proof (rule smallomegaI_filterlim_at_top_norm) from lim show "filterlim (\x. norm (f x / g x)) at_top F" by (rule filterlim_at_infinity_imp_norm_at_top) qed - + lemma smallomegaD_filterlim_at_top_norm: assumes "f \ \[F](g)" assumes "eventually (\x. g x \ 0) F" shows "LIM x F. norm (f x / g x) :> at_top" proof (subst filterlim_at_top_gt, clarify) fix c :: real assume c: "c > 0" - from landau_omega.smallD[OF assms(1) this] assms(2) - show "eventually (\x. norm (f x / g x) \ c) F" + from landau_omega.smallD[OF assms(1) this] assms(2) + show "eventually (\x. norm (f x / g x) \ c) F" by eventually_elim (simp add: field_simps norm_divide) qed - + lemma smallomegaD_filterlim_at_infinity: assumes "f \ \[F](g)" assumes "eventually (\x. g x \ 0) F" shows "LIM x F. f x / g x :> at_infinity" using assms by (intro filterlim_norm_at_top_imp_at_infinity smallomegaD_filterlim_at_top_norm) lemma smallomega_1_conv_filterlim: "f \ \[F](\_. 1) \ filterlim f at_infinity F" by (auto intro: smallomegaI_filterlim_at_infinity dest: smallomegaD_filterlim_at_infinity) lemma smalloI_tendsto: assumes lim: "((\x. f x / g x) \ 0) F" assumes "eventually (\x. g x \ 0) F" shows "f \ o[F](g)" proof (rule landau_o.smallI) fix c :: real assume c_pos: "c > 0" from c_pos and lim have ev: "eventually (\x. norm (f x / g x) < c) F" by (subst (asm) tendsto_iff) (simp add: dist_real_def) with assms(2) show "eventually (\x. (norm (f x)) \ c * (norm (g x))) F" by eventually_elim (simp add: field_simps norm_divide) qed lemma smalloD_tendsto: assumes "f \ o[F](g)" shows "((\x. f x / g x) \ 0) F" unfolding tendsto_iff proof clarify fix e :: real assume e: "e > 0" hence "e/2 > 0" by simp from landau_o.smallD[OF assms this] show "eventually (\x. dist (f x / g x) 0 < e) F" proof eventually_elim fix x assume "(norm (f x)) \ e/2 * (norm (g x))" with e have "dist (f x / g x) 0 \ e/2" by (cases "g x = 0") (simp_all add: dist_real_def norm_divide field_simps) also from e have "... < e" by simp finally show "dist (f x / g x) 0 < e" by simp qed qed lemma bigthetaI_tendsto_norm: assumes c_not_0: "(c::real) \ 0" assumes lim: "((\x. norm (f x / g x)) \ c) F" shows "f \ \[F](g)" proof (rule bigthetaI) from c_not_0 have "\c\ > 0" by simp with lim have "eventually (\x. norm (norm (f x / g x) - c) < \c\) F" by (subst (asm) tendsto_iff) (simp add: dist_real_def) hence g: "eventually (\x. g x \ 0) F" by eventually_elim (auto simp add: field_simps) from lim g show "f \ O[F](g)" by (rule bigoI_tendsto_norm) from c_not_0 and lim show "f \ \[F](g)" by (rule bigomegaI_tendsto_norm) qed lemma bigthetaI_tendsto: assumes c_not_0: "(c::real) \ 0" assumes lim: "((\x. f x / g x) \ c) F" shows "f \ \[F](g)" using assms by (intro bigthetaI_tendsto_norm[OF _ tendsto_norm, of "c"]) simp_all lemma tendsto_add_smallo: assumes "(f1 \ a) F" assumes "f2 \ o[F](f1)" shows "((\x. f1 x + f2 x) \ a) F" proof (subst filterlim_cong[OF refl refl]) - from landau_o.smallD[OF assms(2) zero_less_one] + from landau_o.smallD[OF assms(2) zero_less_one] have "eventually (\x. norm (f2 x) \ norm (f1 x)) F" by simp thus "eventually (\x. f1 x + f2 x = f1 x * (1 + f2 x / f1 x)) F" by eventually_elim (auto simp: field_simps) next from assms(1) show "((\x. f1 x * (1 + f2 x / f1 x)) \ a) F" by (force intro: tendsto_eq_intros smalloD_tendsto[OF assms(2)]) qed lemma tendsto_diff_smallo: shows "(f1 \ a) F \ f2 \ o[F](f1) \ ((\x. f1 x - f2 x) \ a) F" using tendsto_add_smallo[of f1 a F "\x. -f2 x"] by simp lemma tendsto_add_smallo_iff: assumes "f2 \ o[F](f1)" shows "(f1 \ a) F \ ((\x. f1 x + f2 x) \ a) F" proof assume "((\x. f1 x + f2 x) \ a) F" hence "((\x. f1 x + f2 x - f2 x) \ a) F" by (rule tendsto_diff_smallo) (simp add: landau_o.small.plus_absorb2 assms) thus "(f1 \ a) F" by simp qed (rule tendsto_add_smallo[OF _ assms]) lemma tendsto_diff_smallo_iff: shows "f2 \ o[F](f1) \ (f1 \ a) F \ ((\x. f1 x - f2 x) \ a) F" using tendsto_add_smallo_iff[of "\x. -f2 x" F f1 a] by simp lemma tendsto_divide_smallo: assumes "((\x. f1 x / g1 x) \ a) F" assumes "f2 \ o[F](f1)" "g2 \ o[F](g1)" assumes "eventually (\x. g1 x \ 0) F" shows "((\x. (f1 x + f2 x) / (g1 x + g2 x)) \ a) F" (is "(?f \ _) _") proof (subst tendsto_cong) let ?f' = "\x. (f1 x / g1 x) * (1 + f2 x / f1 x) / (1 + g2 x / g1 x)" have "(?f' \ a * (1 + 0) / (1 + 0)) F" - by (rule tendsto_mult tendsto_divide tendsto_add assms tendsto_const + by (rule tendsto_mult tendsto_divide tendsto_add assms tendsto_const smalloD_tendsto[OF assms(2)] smalloD_tendsto[OF assms(3)])+ simp_all thus "(?f' \ a) F" by simp have "(1/2::real) > 0" by simp from landau_o.smallD[OF assms(2) this] landau_o.smallD[OF assms(3) this] have "eventually (\x. norm (f2 x) \ norm (f1 x)/2) F" "eventually (\x. norm (g2 x) \ norm (g1 x)/2) F" by simp_all with assms(4) show "eventually (\x. ?f x = ?f' x) F" proof eventually_elim - fix x assume A: "norm (f2 x) \ norm (f1 x)/2" and + fix x assume A: "norm (f2 x) \ norm (f1 x)/2" and B: "norm (g2 x) \ norm (g1 x)/2" and C: "g1 x \ 0" show "?f x = ?f' x" proof (cases "f1 x = 0") assume D: "f1 x \ 0" from D have "f1 x + f2 x = f1 x * (1 + f2 x/f1 x)" by (simp add: field_simps) moreover from C have "g1 x + g2 x = g1 x * (1 + g2 x/g1 x)" by (simp add: field_simps) ultimately have "?f x = (f1 x * (1 + f2 x/f1 x)) / (g1 x * (1 + g2 x/g1 x))" by (simp only:) also have "... = ?f' x" by simp finally show ?thesis . qed (insert A, simp) qed qed lemma bigo_powr: fixes f :: "'a \ real" assumes "f \ O[F](g)" "p \ 0" shows "(\x. \f x\ powr p) \ O[F](\x. \g x\ powr p)" proof- - from assms(1) guess c by (elim landau_o.bigE landau_omega.bigE IntE) - note c = this - from c(2) assms(2) have "eventually (\x. (norm (f x)) powr p \ (c * (norm (g x))) powr p) F" + from assms(1) obtain c where c: "c > 0" and *: "\\<^sub>F x in F. norm (f x) \ c * norm (g x)" + by (elim landau_o.bigE landau_omega.bigE IntE) + from assms(2) * have "eventually (\x. (norm (f x)) powr p \ (c * norm (g x)) powr p) F" by (auto elim!: eventually_mono intro!: powr_mono2) - thus "(\x. \f x\ powr p) \ O[F](\x. \g x\ powr p)" using c(1) + with c show "(\x. \f x\ powr p) \ O[F](\x. \g x\ powr p)" by (intro bigoI[of _ "c powr p"]) (simp_all add: powr_mult) qed lemma smallo_powr: fixes f :: "'a \ real" assumes "f \ o[F](g)" "p > 0" shows "(\x. \f x\ powr p) \ o[F](\x. \g x\ powr p)" proof (rule landau_o.smallI) fix c :: real assume c: "c > 0" hence "c powr (1/p) > 0" by simp - from landau_o.smallD[OF assms(1) this] + from landau_o.smallD[OF assms(1) this] show "eventually (\x. norm (\f x\ powr p) \ c * norm (\g x\ powr p)) F" proof eventually_elim fix x assume "(norm (f x)) \ c powr (1 / p) * (norm (g x))" with assms(2) have "(norm (f x)) powr p \ (c powr (1 / p) * (norm (g x))) powr p" by (intro powr_mono2) simp_all also from assms(2) c have "... = c * (norm (g x)) powr p" by (simp add: field_simps powr_mult powr_powr) finally show "norm (\f x\ powr p) \ c * norm (\g x\ powr p)" by simp qed qed lemma smallo_powr_nonneg: fixes f :: "'a \ real" assumes "f \ o[F](g)" "p > 0" "eventually (\x. f x \ 0) F" "eventually (\x. g x \ 0) F" shows "(\x. f x powr p) \ o[F](\x. g x powr p)" proof - - from assms(3) have "(\x. f x powr p) \ \[F](\x. \f x\ powr p)" + from assms(3) have "(\x. f x powr p) \ \[F](\x. \f x\ powr p)" by (intro bigthetaI_cong) (auto elim!: eventually_mono) also have "(\x. \f x\ powr p) \ o[F](\x. \g x\ powr p)" by (intro smallo_powr) fact+ also from assms(4) have "(\x. \g x\ powr p) \ \[F](\x. g x powr p)" by (intro bigthetaI_cong) (auto elim!: eventually_mono) finally show ?thesis . qed lemma bigtheta_powr: fixes f :: "'a \ real" shows "f \ \[F](g) \ (\x. \f x\ powr p) \ \[F](\x. \g x\ powr p)" apply (cases "p < 0") apply (subst bigtheta_inverse[symmetric], subst (1 2) powr_minus[symmetric]) unfolding bigtheta_def apply (auto simp: bigomega_iff_bigo intro!: bigo_powr) done lemma bigo_powr_nonneg: fixes f :: "'a \ real" assumes "f \ O[F](g)" "p \ 0" "eventually (\x. f x \ 0) F" "eventually (\x. g x \ 0) F" shows "(\x. f x powr p) \ O[F](\x. g x powr p)" proof - - from assms(3) have "(\x. f x powr p) \ \[F](\x. \f x\ powr p)" + from assms(3) have "(\x. f x powr p) \ \[F](\x. \f x\ powr p)" by (intro bigthetaI_cong) (auto elim!: eventually_mono) also have "(\x. \f x\ powr p) \ O[F](\x. \g x\ powr p)" by (intro bigo_powr) fact+ also from assms(4) have "(\x. \g x\ powr p) \ \[F](\x. g x powr p)" by (intro bigthetaI_cong) (auto elim!: eventually_mono) finally show ?thesis . qed lemma zero_in_smallo [simp]: "(\_. 0) \ o[F](f)" by (intro landau_o.smallI) simp_all lemma zero_in_bigo [simp]: "(\_. 0) \ O[F](f)" by (intro landau_o.bigI[of 1]) simp_all lemma in_bigomega_zero [simp]: "f \ \[F](\x. 0)" by (rule landau_omega.bigI[of 1]) simp_all lemma in_smallomega_zero [simp]: "f \ \[F](\x. 0)" by (simp add: smallomega_iff_smallo) lemma in_smallo_zero_iff [simp]: "f \ o[F](\_. 0) \ eventually (\x. f x = 0) F" proof assume "f \ o[F](\_. 0)" from landau_o.smallD[OF this, of 1] show "eventually (\x. f x = 0) F" by simp next assume "eventually (\x. f x = 0) F" hence "\c>0. eventually (\x. (norm (f x)) \ c * \0\) F" by simp thus "f \ o[F](\_. 0)" unfolding smallo_def by simp qed lemma in_bigo_zero_iff [simp]: "f \ O[F](\_. 0) \ eventually (\x. f x = 0) F" proof assume "f \ O[F](\_. 0)" thus "eventually (\x. f x = 0) F" by (elim landau_o.bigE) simp next assume "eventually (\x. f x = 0) F" hence "eventually (\x. (norm (f x)) \ 1 * \0\) F" by simp thus "f \ O[F](\_. 0)" by (intro landau_o.bigI[of 1]) simp_all qed lemma zero_in_smallomega_iff [simp]: "(\_. 0) \ \[F](f) \ eventually (\x. f x = 0) F" by (simp add: smallomega_iff_smallo) lemma zero_in_bigomega_iff [simp]: "(\_. 0) \ \[F](f) \ eventually (\x. f x = 0) F" by (simp add: bigomega_iff_bigo) lemma zero_in_bigtheta_iff [simp]: "(\_. 0) \ \[F](f) \ eventually (\x. f x = 0) F" unfolding bigtheta_def by simp lemma in_bigtheta_zero_iff [simp]: "f \ \[F](\x. 0) \ eventually (\x. f x = 0) F" unfolding bigtheta_def by simp lemma cmult_in_bigo_iff [simp]: "(\x. c * f x) \ O[F](g) \ c = 0 \ f \ O[F](g)" and cmult_in_bigo_iff' [simp]: "(\x. f x * c) \ O[F](g) \ c = 0 \ f \ O[F](g)" and cmult_in_smallo_iff [simp]: "(\x. c * f x) \ o[F](g) \ c = 0 \ f \ o[F](g)" and cmult_in_smallo_iff' [simp]: "(\x. f x * c) \ o[F](g) \ c = 0 \ f \ o[F](g)" by (cases "c = 0", simp, simp)+ lemma bigo_const [simp]: "(\_. c) \ O[F](\_. 1)" by (rule bigoI[of _ "norm c"]) simp lemma bigo_const_iff [simp]: "(\_. c1) \ O[F](\_. c2) \ F = bot \ c1 = 0 \ c2 \ 0" by (cases "c1 = 0"; cases "c2 = 0") (auto simp: bigo_def eventually_False intro: exI[of _ 1] exI[of _ "norm c1 / norm c2"]) lemma bigomega_const_iff [simp]: "(\_. c1) \ \[F](\_. c2) \ F = bot \ c1 \ 0 \ c2 = 0" by (cases "c1 = 0"; cases "c2 = 0") - (auto simp: bigomega_def eventually_False mult_le_0_iff + (auto simp: bigomega_def eventually_False mult_le_0_iff intro: exI[of _ 1] exI[of _ "norm c1 / norm c2"]) lemma smallo_real_nat_transfer: "(f :: real \ real) \ o(g) \ (\x::nat. f (real x)) \ o(\x. g (real x))" by (rule landau_o.small.compose[OF _ filterlim_real_sequentially]) lemma bigo_real_nat_transfer: "(f :: real \ real) \ O(g) \ (\x::nat. f (real x)) \ O(\x. g (real x))" by (rule landau_o.big.compose[OF _ filterlim_real_sequentially]) lemma smallomega_real_nat_transfer: "(f :: real \ real) \ \(g) \ (\x::nat. f (real x)) \ \(\x. g (real x))" by (rule landau_omega.small.compose[OF _ filterlim_real_sequentially]) lemma bigomega_real_nat_transfer: "(f :: real \ real) \ \(g) \ (\x::nat. f (real x)) \ \(\x. g (real x))" by (rule landau_omega.big.compose[OF _ filterlim_real_sequentially]) lemma bigtheta_real_nat_transfer: "(f :: real \ real) \ \(g) \ (\x::nat. f (real x)) \ \(\x. g (real x))" unfolding bigtheta_def using bigo_real_nat_transfer bigomega_real_nat_transfer by blast -lemmas landau_real_nat_transfer [intro] = - bigo_real_nat_transfer smallo_real_nat_transfer bigomega_real_nat_transfer +lemmas landau_real_nat_transfer [intro] = + bigo_real_nat_transfer smallo_real_nat_transfer bigomega_real_nat_transfer smallomega_real_nat_transfer bigtheta_real_nat_transfer lemma landau_symbol_if_at_top_eq [simp]: assumes "landau_symbol L L' Lr" shows "L at_top (\x::'a::linordered_semidom. if x = a then f x else g x) = L at_top (g)" apply (rule landau_symbol.cong[OF assms]) using less_add_one[of a] apply (auto intro: eventually_mono eventually_ge_at_top[of "a + 1"]) done lemmas landau_symbols_if_at_top_eq [simp] = landau_symbols[THEN landau_symbol_if_at_top_eq] lemma sum_in_smallo: assumes "f \ o[F](h)" "g \ o[F](h)" shows "(\x. f x + g x) \ o[F](h)" "(\x. f x - g x) \ o[F](h)" -proof- - { - fix f g assume fg: "f \ o[F](h)" "g \ o[F](h)" - have "(\x. f x + g x) \ o[F](h)" - proof (rule landau_o.smallI) - fix c :: real assume "c > 0" - hence "c/2 > 0" by simp - from fg[THEN landau_o.smallD[OF _ this]] - show "eventually (\x. norm (f x + g x) \ c * (norm (h x))) F" - by eventually_elim (auto intro: order.trans[OF norm_triangle_ineq]) - qed - } +proof - + have "(\x. f x + g x) \ o[F](h)" if "f \ o[F](h)" "g \ o[F](h)" for f g + proof (rule landau_o.smallI) + fix c :: real assume "c > 0" + hence "c/2 > 0" by simp + from that[THEN landau_o.smallD[OF _ this]] + show "eventually (\x. norm (f x + g x) \ c * (norm (h x))) F" + by eventually_elim (auto intro: order.trans[OF norm_triangle_ineq]) + qed from this[of f g] this[of f "\x. -g x"] assms show "(\x. f x + g x) \ o[F](h)" "(\x. f x - g x) \ o[F](h)" by simp_all qed lemma big_sum_in_smallo: assumes "\x. x \ A \ f x \ o[F](g)" shows "(\x. sum (\y. f y x) A) \ o[F](g)" using assms by (induction A rule: infinite_finite_induct) (auto intro: sum_in_smallo) lemma sum_in_bigo: assumes "f \ O[F](h)" "g \ O[F](h)" shows "(\x. f x + g x) \ O[F](h)" "(\x. f x - g x) \ O[F](h)" -proof- - { - fix f g assume fg: "f \ O[F](h)" "g \ O[F](h)" - from fg(1) guess c1 by (elim landau_o.bigE) note c1 = this - from fg(2) guess c2 by (elim landau_o.bigE) note c2 = this - from c1(2) c2(2) have "eventually (\x. norm (f x + g x) \ (c1 + c2) * (norm (h x))) F" +proof - + have "(\x. f x + g x) \ O[F](h)" if "f \ O[F](h)" "g \ O[F](h)" for f g + proof - + from that obtain c1 c2 where *: "c1 > 0" "c2 > 0" + and **: "\\<^sub>F x in F. norm (f x) \ c1 * norm (h x)" + "\\<^sub>F x in F. norm (g x) \ c2 * norm (h x)" + by (elim landau_o.bigE) + from ** have "eventually (\x. norm (f x + g x) \ (c1 + c2) * (norm (h x))) F" by eventually_elim (auto simp: algebra_simps intro: order.trans[OF norm_triangle_ineq]) - hence "(\x. f x + g x) \ O[F](h)" by (rule bigoI) - } - from this[of f g] this[of f "\x. -g x"] assms - show "(\x. f x + g x) \ O[F](h)" "(\x. f x - g x) \ O[F](h)" by simp_all + then show ?thesis by (rule bigoI) + qed + from assms this[of f g] this[of f "\x. - g x"] + show "(\x. f x + g x) \ O[F](h)" "(\x. f x - g x) \ O[F](h)" by simp_all qed lemma big_sum_in_bigo: assumes "\x. x \ A \ f x \ O[F](g)" shows "(\x. sum (\y. f y x) A) \ O[F](g)" using assms by (induction A rule: infinite_finite_induct) (auto intro: sum_in_bigo) context landau_symbol begin lemma mult_cancel_left: assumes "f1 \ \[F](g1)" and "eventually (\x. g1 x \ 0) F" notes [trans] = bigtheta_trans1 bigtheta_trans2 shows "(\x. f1 x * f2 x) \ L F (\x. g1 x * g2 x) \ f2 \ L F (g2)" proof assume A: "(\x. f1 x * f2 x) \ L F (\x. g1 x * g2 x)" from assms have nz: "eventually (\x. f1 x \ 0) F" by (simp add: eventually_nonzero_bigtheta) hence "f2 \ \[F](\x. f1 x * f2 x / f1 x)" by (intro bigthetaI_cong) (auto elim!: eventually_mono) - also from A assms nz have "(\x. f1 x * f2 x / f1 x) \ L F (\x. g1 x * g2 x / f1 x)" + also from A assms nz have "(\x. f1 x * f2 x / f1 x) \ L F (\x. g1 x * g2 x / f1 x)" by (intro divide_right) simp_all also from assms nz have "(\x. g1 x * g2 x / f1 x) \ \[F](\x. g1 x * g2 x / g1 x)" by (intro landau_theta.mult landau_theta.divide) (simp_all add: bigtheta_sym) also from assms have "(\x. g1 x * g2 x / g1 x) \ \[F](g2)" by (intro bigthetaI_cong) (auto elim!: eventually_mono) finally show "f2 \ L F (g2)" . next assume "f2 \ L F (g2)" hence "(\x. f1 x * f2 x) \ L F (\x. f1 x * g2 x)" by (rule mult_left) also have "(\x. f1 x * g2 x) \ \[F](\x. g1 x * g2 x)" by (intro landau_theta.mult_right assms) finally show "(\x. f1 x * f2 x) \ L F (\x. g1 x * g2 x)" . qed lemma mult_cancel_right: assumes "f2 \ \[F](g2)" and "eventually (\x. g2 x \ 0) F" shows "(\x. f1 x * f2 x) \ L F (\x. g1 x * g2 x) \ f1 \ L F (g1)" by (subst (1 2) mult.commute) (rule mult_cancel_left[OF assms]) lemma divide_cancel_right: assumes "f2 \ \[F](g2)" and "eventually (\x. g2 x \ 0) F" shows "(\x. f1 x / f2 x) \ L F (\x. g1 x / g2 x) \ f1 \ L F (g1)" by (subst (1 2) divide_inverse, intro mult_cancel_right bigtheta_inverse) (simp_all add: assms) lemma divide_cancel_left: assumes "f1 \ \[F](g1)" and "eventually (\x. g1 x \ 0) F" - shows "(\x. f1 x / f2 x) \ L F (\x. g1 x / g2 x) \ + shows "(\x. f1 x / f2 x) \ L F (\x. g1 x / g2 x) \ (\x. inverse (f2 x)) \ L F (\x. inverse (g2 x))" by (simp only: divide_inverse mult_cancel_left[OF assms]) end lemma powr_smallo_iff: assumes "filterlim g at_top F" "F \ bot" shows "(\x. g x powr p :: real) \ o[F](\x. g x powr q) \ p < q" proof- from assms have "eventually (\x. g x \ 1) F" by (force simp: filterlim_at_top) hence A: "eventually (\x. g x \ 0) F" by eventually_elim simp have B: "(\x. g x powr q) \ O[F](\x. g x powr p) \ (\x. g x powr p) \ o[F](\x. g x powr q)" proof assume "(\x. g x powr q) \ O[F](\x. g x powr p)" "(\x. g x powr p) \ o[F](\x. g x powr q)" from landau_o.big_small_asymmetric[OF this] have "eventually (\x. g x = 0) F" by simp with A have "eventually (\_::'a. False) F" by eventually_elim simp thus False by (simp add: eventually_False assms) qed show ?thesis proof (cases p q rule: linorder_cases) assume "p < q" hence "(\x. g x powr p) \ o[F](\x. g x powr q)" using assms A by (auto intro!: smalloI_tendsto tendsto_neg_powr simp flip: powr_diff) with \p < q\ show ?thesis by auto next assume "p = q" hence "(\x. g x powr q) \ O[F](\x. g x powr p)" by (auto intro!: bigthetaD1) with B \p = q\ show ?thesis by auto next assume "p > q" hence "(\x. g x powr q) \ O[F](\x. g x powr p)" using assms A by (auto intro!: smalloI_tendsto tendsto_neg_powr landau_o.small_imp_big simp flip: powr_diff) with B \p > q\ show ?thesis by auto qed qed lemma powr_bigo_iff: assumes "filterlim g at_top F" "F \ bot" shows "(\x. g x powr p :: real) \ O[F](\x. g x powr q) \ p \ q" proof- from assms have "eventually (\x. g x \ 1) F" by (force simp: filterlim_at_top) hence A: "eventually (\x. g x \ 0) F" by eventually_elim simp have B: "(\x. g x powr q) \ o[F](\x. g x powr p) \ (\x. g x powr p) \ O[F](\x. g x powr q)" proof assume "(\x. g x powr q) \ o[F](\x. g x powr p)" "(\x. g x powr p) \ O[F](\x. g x powr q)" from landau_o.small_big_asymmetric[OF this] have "eventually (\x. g x = 0) F" by simp with A have "eventually (\_::'a. False) F" by eventually_elim simp thus False by (simp add: eventually_False assms) qed show ?thesis proof (cases p q rule: linorder_cases) assume "p < q" hence "(\x. g x powr p) \ o[F](\x. g x powr q)" using assms A by (auto intro!: smalloI_tendsto tendsto_neg_powr simp flip: powr_diff) with \p < q\ show ?thesis by (auto intro: landau_o.small_imp_big) next assume "p = q" hence "(\x. g x powr q) \ O[F](\x. g x powr p)" by (auto intro!: bigthetaD1) with B \p = q\ show ?thesis by auto next assume "p > q" hence "(\x. g x powr q) \ o[F](\x. g x powr p)" using assms A by (auto intro!: smalloI_tendsto tendsto_neg_powr simp flip: powr_diff) with B \p > q\ show ?thesis by (auto intro: landau_o.small_imp_big) qed qed -lemma powr_bigtheta_iff: +lemma powr_bigtheta_iff: assumes "filterlim g at_top F" "F \ bot" shows "(\x. g x powr p :: real) \ \[F](\x. g x powr q) \ p = q" using assms unfolding bigtheta_def by (auto simp: bigomega_iff_bigo powr_bigo_iff) subsection \Flatness of real functions\ text \ Given two real-valued functions $f$ and $g$, we say that $f$ is flatter than $g$ if any power of $f(x)$ is asymptotically dominated by any positive power of $g(x)$. This is a useful notion since, given two products of powers of functions sorted by flatness, we can compare them asymptotically by simply comparing the exponent lists lexicographically. A simple sufficient criterion for flatness it that $\ln f(x) \in o(\ln g(x))$, which we show now. \ lemma ln_smallo_imp_flat: fixes f g :: "real \ real" assumes lim_f: "filterlim f at_top at_top" assumes lim_g: "filterlim g at_top at_top" assumes ln_o_ln: "(\x. ln (f x)) \ o(\x. ln (g x))" assumes q: "q > 0" shows "(\x. f x powr p) \ o(\x. g x powr q)" proof (rule smalloI_tendsto) - from lim_f have "eventually (\x. f x > 0) at_top" + from lim_f have "eventually (\x. f x > 0) at_top" by (simp add: filterlim_at_top_dense) hence f_nz: "eventually (\x. f x \ 0) at_top" by eventually_elim simp - + from lim_g have g_gt_1: "eventually (\x. g x > 1) at_top" by (simp add: filterlim_at_top_dense) hence g_nz: "eventually (\x. g x \ 0) at_top" by eventually_elim simp thus "eventually (\x. g x powr q \ 0) at_top" by eventually_elim simp - - have eq: "eventually (\x. q * (p/q * (ln (f x) / ln (g x)) - 1) * ln (g x) = + + have eq: "eventually (\x. q * (p/q * (ln (f x) / ln (g x)) - 1) * ln (g x) = p * ln (f x) - q * ln (g x)) at_top" using g_gt_1 by eventually_elim (insert q, simp_all add: field_simps) have "filterlim (\x. q * (p/q * (ln (f x) / ln (g x)) - 1) * ln (g x)) at_bot at_top" by (insert q) (rule filterlim_tendsto_neg_mult_at_bot tendsto_mult tendsto_const tendsto_diff smalloD_tendsto[OF ln_o_ln] lim_g filterlim_compose[OF ln_at_top] | simp)+ hence "filterlim (\x. p * ln (f x) - q * ln (g x)) at_bot at_top" by (subst (asm) filterlim_cong[OF refl refl eq]) hence *: "((\x. exp (p * ln (f x) - q * ln (g x))) \ 0) at_top" by (rule filterlim_compose[OF exp_at_bot]) have eq: "eventually (\x. exp (p * ln (f x) - q * ln (g x)) = f x powr p / g x powr q) at_top" using f_nz g_nz by eventually_elim (simp add: powr_def exp_diff) show "((\x. f x powr p / g x powr q) \ 0) at_top" using * by (subst (asm) filterlim_cong[OF refl refl eq]) qed lemma ln_smallo_imp_flat': fixes f g :: "real \ real" assumes lim_f: "filterlim f at_top at_top" assumes lim_g: "filterlim g at_top at_top" assumes ln_o_ln: "(\x. ln (f x)) \ o(\x. ln (g x))" assumes q: "q < 0" shows "(\x. g x powr q) \ o(\x. f x powr p)" proof - from lim_f lim_g have "eventually (\x. f x > 0) at_top" "eventually (\x. g x > 0) at_top" by (simp_all add: filterlim_at_top_dense) hence "eventually (\x. f x \ 0) at_top" "eventually (\x. g x \ 0) at_top" by (auto elim: eventually_mono) moreover from assms have "(\x. f x powr -p) \ o(\x. g x powr -q)" by (intro ln_smallo_imp_flat assms) simp_all ultimately show ?thesis unfolding powr_minus by (simp add: landau_o.small.inverse_cancel) qed subsection \Asymptotic Equivalence\ named_theorems asymp_equiv_intros named_theorems asymp_equiv_simps definition asymp_equiv :: "('a \ ('b :: real_normed_field)) \ 'a filter \ ('a \ 'b) \ bool" (\_ \[_] _\ [51, 10, 51] 50) where "f \[F] g \ ((\x. if f x = 0 \ g x = 0 then 1 else f x / g x) \ 1) F" abbreviation (input) asymp_equiv_at_top where "asymp_equiv_at_top f g \ f \[at_top] g" bundle asymp_equiv_notation begin -notation asymp_equiv_at_top (infix \\\ 50) +notation asymp_equiv_at_top (infix \\\ 50) end lemma asymp_equivI: "((\x. if f x = 0 \ g x = 0 then 1 else f x / g x) \ 1) F \ f \[F] g" by (simp add: asymp_equiv_def) lemma asymp_equivD: "f \[F] g \ ((\x. if f x = 0 \ g x = 0 then 1 else f x / g x) \ 1) F" by (simp add: asymp_equiv_def) lemma asymp_equiv_filtermap_iff: "f \[filtermap h F] g \ (\x. f (h x)) \[F] (\x. g (h x))" by (simp add: asymp_equiv_def filterlim_filtermap) lemma asymp_equiv_refl [simp, asymp_equiv_intros]: "f \[F] f" proof (intro asymp_equivI) have "eventually (\x. 1 = (if f x = 0 \ f x = 0 then 1 else f x / f x)) F" by (intro always_eventually) simp moreover have "((\_. 1) \ 1) F" by simp ultimately show "((\x. if f x = 0 \ f x = 0 then 1 else f x / f x) \ 1) F" by (simp add: tendsto_eventually) qed -lemma asymp_equiv_symI: +lemma asymp_equiv_symI: assumes "f \[F] g" shows "g \[F] f" using tendsto_inverse[OF asymp_equivD[OF assms]] by (auto intro!: asymp_equivI simp: if_distrib conj_commute cong: if_cong) lemma asymp_equiv_sym: "f \[F] g \ g \[F] f" by (blast intro: asymp_equiv_symI) -lemma asymp_equivI': +lemma asymp_equivI': assumes "((\x. f x / g x) \ 1) F" shows "f \[F] g" proof (cases "F = bot") case False have "eventually (\x. f x \ 0) F" proof (rule ccontr) assume "\eventually (\x. f x \ 0) F" hence "frequently (\x. f x = 0) F" by (simp add: frequently_def) hence "frequently (\x. f x / g x = 0) F" by (auto elim!: frequently_elim1) from limit_frequently_eq[OF False this assms] show False by simp_all qed hence "eventually (\x. f x / g x = (if f x = 0 \ g x = 0 then 1 else f x / g x)) F" by eventually_elim simp - with assms show "f \[F] g" unfolding asymp_equiv_def + with assms show "f \[F] g" unfolding asymp_equiv_def by (rule Lim_transform_eventually) qed (simp_all add: asymp_equiv_def) lemma asymp_equiv_cong: assumes "eventually (\x. f1 x = f2 x) F" "eventually (\x. g1 x = g2 x) F" shows "f1 \[F] g1 \ f2 \[F] g2" unfolding asymp_equiv_def proof (rule tendsto_cong, goal_cases) case 1 from assms show ?case by eventually_elim simp qed lemma asymp_equiv_eventually_zeros: fixes f g :: "'a \ 'b :: real_normed_field" assumes "f \[F] g" shows "eventually (\x. f x = 0 \ g x = 0) F" proof - let ?h = "\x. if f x = 0 \ g x = 0 then 1 else f x / g x" have "eventually (\x. x \ 0) (nhds (1::'b))" by (rule t1_space_nhds) auto hence "eventually (\x. x \ 0) (filtermap ?h F)" using assms unfolding asymp_equiv_def filterlim_def by (rule filter_leD [rotated]) hence "eventually (\x. ?h x \ 0) F" by (simp add: eventually_filtermap) thus ?thesis by eventually_elim (auto split: if_splits) qed lemma asymp_equiv_transfer: assumes "f1 \[F] g1" "eventually (\x. f1 x = f2 x) F" "eventually (\x. g1 x = g2 x) F" shows "f2 \[F] g2" using assms(1) asymp_equiv_cong[OF assms(2,3)] by simp lemma asymp_equiv_transfer_trans [trans]: assumes "(\x. f x (h1 x)) \[F] (\x. g x (h1 x))" assumes "eventually (\x. h1 x = h2 x) F" shows "(\x. f x (h2 x)) \[F] (\x. g x (h2 x))" by (rule asymp_equiv_transfer[OF assms(1)]) (insert assms(2), auto elim!: eventually_mono) lemma asymp_equiv_trans [trans]: fixes f g h assumes "f \[F] g" "g \[F] h" shows "f \[F] h" proof - let ?T = "\f g x. if f x = 0 \ g x = 0 then 1 else f x / g x" - from tendsto_mult[OF assms[THEN asymp_equivD]] + from tendsto_mult[OF assms[THEN asymp_equivD]] have "((\x. ?T f g x * ?T g h x) \ 1) F" by simp moreover from assms[THEN asymp_equiv_eventually_zeros] have "eventually (\x. ?T f g x * ?T g h x = ?T f h x) F" by eventually_elim simp ultimately show ?thesis unfolding asymp_equiv_def by (rule Lim_transform_eventually) qed lemma asymp_equiv_trans_lift1 [trans]: assumes "a \[F] f b" "b \[F] c" "\c d. c \[F] d \ f c \[F] f d" shows "a \[F] f c" using assms by (blast intro: asymp_equiv_trans) lemma asymp_equiv_trans_lift2 [trans]: assumes "f a \[F] b" "a \[F] c" "\c d. c \[F] d \ f c \[F] f d" shows "f c \[F] b" using asymp_equiv_symI[OF assms(3)[OF assms(2)]] assms(1) by (blast intro: asymp_equiv_trans) lemma asymp_equivD_const: assumes "f \[F] (\_. c)" shows "(f \ c) F" proof (cases "c = 0") case False with tendsto_mult_right[OF asymp_equivD[OF assms], of c] show ?thesis by simp next case True with asymp_equiv_eventually_zeros[OF assms] show ?thesis by (simp add: tendsto_eventually) qed lemma asymp_equiv_refl_ev: assumes "eventually (\x. f x = g x) F" shows "f \[F] g" by (intro asymp_equivI tendsto_eventually) (insert assms, auto elim!: eventually_mono) lemma asymp_equiv_sandwich: fixes f g h :: "'a \ 'b :: {real_normed_field, order_topology, linordered_field}" assumes "eventually (\x. f x \ 0) F" assumes "eventually (\x. f x \ g x) F" assumes "eventually (\x. g x \ h x) F" assumes "f \[F] h" shows "g \[F] f" "g \[F] h" proof - show "g \[F] f" proof (rule asymp_equivI, rule tendsto_sandwich) from assms(1-3) asymp_equiv_eventually_zeros[OF assms(4)] show "eventually (\n. (if h n = 0 \ f n = 0 then 1 else h n / f n) \ (if g n = 0 \ f n = 0 then 1 else g n / f n)) F" by eventually_elim (auto intro!: divide_right_mono) from assms(1-3) asymp_equiv_eventually_zeros[OF assms(4)] show "eventually (\n. 1 \ (if g n = 0 \ f n = 0 then 1 else g n / f n)) F" by eventually_elim (auto intro!: divide_right_mono) qed (insert asymp_equiv_symI[OF assms(4)], simp_all add: asymp_equiv_def) also note \f \[F] h\ finally show "g \[F] h" . qed lemma asymp_equiv_imp_eventually_same_sign: fixes f g :: "real \ real" assumes "f \[F] g" shows "eventually (\x. sgn (f x) = sgn (g x)) F" proof - from assms have "((\x. sgn (if f x = 0 \ g x = 0 then 1 else f x / g x)) \ sgn 1) F" unfolding asymp_equiv_def by (rule tendsto_sgn) simp_all from order_tendstoD(1)[OF this, of "1/2"] have "eventually (\x. sgn (if f x = 0 \ g x = 0 then 1 else f x / g x) > 1/2) F" by simp thus "eventually (\x. sgn (f x) = sgn (g x)) F" proof eventually_elim case (elim x) thus ?case - by (cases "f x" "0 :: real" rule: linorder_cases; + by (cases "f x" "0 :: real" rule: linorder_cases; cases "g x" "0 :: real" rule: linorder_cases) simp_all qed qed lemma fixes f g :: "_ \ real" assumes "f \[F] g" shows asymp_equiv_eventually_same_sign: "eventually (\x. sgn (f x) = sgn (g x)) F" (is ?th1) and asymp_equiv_eventually_neg_iff: "eventually (\x. f x < 0 \ g x < 0) F" (is ?th2) and asymp_equiv_eventually_pos_iff: "eventually (\x. f x > 0 \ g x > 0) F" (is ?th3) proof - from assms have "filterlim (\x. if f x = 0 \ g x = 0 then 1 else f x / g x) (nhds 1) F" by (rule asymp_equivD) from order_tendstoD(1)[OF this zero_less_one] show ?th1 ?th2 ?th3 by (eventually_elim; force simp: sgn_if field_split_simps split: if_splits)+ qed lemma asymp_equiv_tendsto_transfer: assumes "f \[F] g" and "(f \ c) F" shows "(g \ c) F" proof - let ?h = "\x. (if g x = 0 \ f x = 0 then 1 else g x / f x) * f x" from assms(1) have "g \[F] f" by (rule asymp_equiv_symI) hence "filterlim (\x. if g x = 0 \ f x = 0 then 1 else g x / f x) (nhds 1) F" by (rule asymp_equivD) from tendsto_mult[OF this assms(2)] have "(?h \ c) F" by simp - moreover + moreover have "eventually (\x. ?h x = g x) F" using asymp_equiv_eventually_zeros[OF assms(1)] by eventually_elim simp ultimately show ?thesis by (rule Lim_transform_eventually) qed lemma tendsto_asymp_equiv_cong: assumes "f \[F] g" shows "(f \ c) F \ (g \ c) F" proof - - { - fix f g :: "'a \ 'b" - assume *: "f \[F] g" "(g \ c) F" - have "((\x. g x * (if f x = 0 \ g x = 0 then 1 else f x / g x)) \ c * 1) F" - by (intro tendsto_intros asymp_equivD *) - moreover + have "(f \ c * 1) F" if fg: "f \[F] g" and "(g \ c) F" for f g :: "'a \ 'b" + proof - + from that have *: "((\x. g x * (if f x = 0 \ g x = 0 then 1 else f x / g x)) \ c * 1) F" + by (intro tendsto_intros asymp_equivD) have "eventually (\x. g x * (if f x = 0 \ g x = 0 then 1 else f x / g x) = f x) F" - using asymp_equiv_eventually_zeros[OF *(1)] by eventually_elim simp - ultimately have "(f \ c * 1) F" - by (rule Lim_transform_eventually) - } + using asymp_equiv_eventually_zeros[OF fg] by eventually_elim simp + with * show ?thesis by (rule Lim_transform_eventually) + qed from this[of f g] this[of g f] assms show ?thesis by (auto simp: asymp_equiv_sym) qed lemma smallo_imp_eventually_sgn: fixes f g :: "real \ real" assumes "g \ o(f)" shows "eventually (\x. sgn (f x + g x) = sgn (f x)) at_top" proof - have "0 < (1/2 :: real)" by simp - from landau_o.smallD[OF assms, OF this] + from landau_o.smallD[OF assms, OF this] have "eventually (\x. \g x\ \ 1/2 * \f x\) at_top" by simp thus ?thesis proof eventually_elim case (elim x) thus ?case - by (cases "f x" "0::real" rule: linorder_cases; + by (cases "f x" "0::real" rule: linorder_cases; cases "f x + g x" "0::real" rule: linorder_cases) simp_all qed qed context begin private lemma asymp_equiv_add_rightI: assumes "f \[F] g" "h \ o[F](g)" shows "(\x. f x + h x) \[F] g" proof - let ?T = "\f g x. if f x = 0 \ g x = 0 then 1 else f x / g x" from landau_o.smallD[OF assms(2) zero_less_one] have ev: "eventually (\x. g x = 0 \ h x = 0) F" by eventually_elim auto have "(\x. f x + h x) \[F] g \ ((\x. ?T f g x + h x / g x) \ 1) F" unfolding asymp_equiv_def using ev by (intro tendsto_cong) (auto elim!: eventually_mono simp: field_split_simps) also have "\ \ ((\x. ?T f g x + h x / g x) \ 1 + 0) F" by simp also have \ by (intro tendsto_intros asymp_equivD assms smalloD_tendsto) finally show "(\x. f x + h x) \[F] g" . qed lemma asymp_equiv_add_right [asymp_equiv_simps]: assumes "h \ o[F](g)" shows "(\x. f x + h x) \[F] g \ f \[F] g" proof assume "(\x. f x + h x) \[F] g" from asymp_equiv_add_rightI[OF this, of "\x. -h x"] assms show "f \[F] g" by simp qed (simp_all add: asymp_equiv_add_rightI assms) end -lemma asymp_equiv_add_left [asymp_equiv_simps]: +lemma asymp_equiv_add_left [asymp_equiv_simps]: assumes "h \ o[F](g)" shows "(\x. h x + f x) \[F] g \ f \[F] g" using asymp_equiv_add_right[OF assms] by (simp add: add.commute) lemma asymp_equiv_add_right' [asymp_equiv_simps]: assumes "h \ o[F](g)" shows "g \[F] (\x. f x + h x) \ g \[F] f" using asymp_equiv_add_right[OF assms] by (simp add: asymp_equiv_sym) - + lemma asymp_equiv_add_left' [asymp_equiv_simps]: assumes "h \ o[F](g)" shows "g \[F] (\x. h x + f x) \ g \[F] f" using asymp_equiv_add_left[OF assms] by (simp add: asymp_equiv_sym) lemma smallo_imp_asymp_equiv: assumes "(\x. f x - g x) \ o[F](g)" shows "f \[F] g" proof - from assms have "(\x. f x - g x + g x) \[F] g" by (subst asymp_equiv_add_left) simp_all thus ?thesis by simp qed lemma asymp_equiv_uminus [asymp_equiv_intros]: "f \[F] g \ (\x. -f x) \[F] (\x. -g x)" by (simp add: asymp_equiv_def cong: if_cong) lemma asymp_equiv_uminus_iff [asymp_equiv_simps]: "(\x. -f x) \[F] g \ f \[F] (\x. -g x)" by (simp add: asymp_equiv_def cong: if_cong) lemma asymp_equiv_mult [asymp_equiv_intros]: fixes f1 f2 g1 g2 :: "'a \ 'b :: real_normed_field" assumes "f1 \[F] g1" "f2 \[F] g2" shows "(\x. f1 x * f2 x) \[F] (\x. g1 x * g2 x)" proof - let ?T = "\f g x. if f x = 0 \ g x = 0 then 1 else f x / g x" let ?S = "\x. (if f1 x = 0 \ g1 x = 0 then 1 - ?T f2 g2 x else if f2 x = 0 \ g2 x = 0 then 1 - ?T f1 g1 x else 0)" let ?S' = "\x. ?T (\x. f1 x * f2 x) (\x. g1 x * g2 x) x - ?T f1 g1 x * ?T f2 g2 x" + have A: "((\x. 1 - ?T f g x) \ 0) F" if "f \[F] g" for f g :: "'a \ 'b" + by (rule tendsto_eq_intros refl asymp_equivD[OF that])+ simp_all + + from assms have *: "((\x. ?T f1 g1 x * ?T f2 g2 x) \ 1 * 1) F" + by (intro tendsto_mult asymp_equivD) { - fix f g :: "'a \ 'b" assume "f \[F] g" - have "((\x. 1 - ?T f g x) \ 0) F" - by (rule tendsto_eq_intros refl asymp_equivD[OF \f \[F] g\])+ simp_all - } note A = this - - from assms have "((\x. ?T f1 g1 x * ?T f2 g2 x) \ 1 * 1) F" - by (intro tendsto_mult asymp_equivD) - moreover { have "(?S \ 0) F" by (intro filterlim_If assms[THEN A, THEN tendsto_mono[rotated]]) (auto intro: le_infI1 le_infI2) moreover have "eventually (\x. ?S x = ?S' x) F" using assms[THEN asymp_equiv_eventually_zeros] by eventually_elim auto ultimately have "(?S' \ 0) F" by (rule Lim_transform_eventually) } - ultimately have "(?T (\x. f1 x * f2 x) (\x. g1 x * g2 x) \ 1 * 1) F" + with * have "(?T (\x. f1 x * f2 x) (\x. g1 x * g2 x) \ 1 * 1) F" by (rule Lim_transform) - thus ?thesis by (simp add: asymp_equiv_def) + then show ?thesis by (simp add: asymp_equiv_def) qed lemma asymp_equiv_power [asymp_equiv_intros]: "f \[F] g \ (\x. f x ^ n) \[F] (\x. g x ^ n)" by (induction n) (simp_all add: asymp_equiv_mult) lemma asymp_equiv_inverse [asymp_equiv_intros]: assumes "f \[F] g" shows "(\x. inverse (f x)) \[F] (\x. inverse (g x))" proof - from tendsto_inverse[OF asymp_equivD[OF assms]] have "((\x. if f x = 0 \ g x = 0 then 1 else g x / f x) \ 1) F" by (simp add: if_distrib cong: if_cong) also have "(\x. if f x = 0 \ g x = 0 then 1 else g x / f x) = (\x. if f x = 0 \ g x = 0 then 1 else inverse (f x) / inverse (g x))" by (intro ext) (simp add: field_simps) finally show ?thesis by (simp add: asymp_equiv_def) qed lemma asymp_equiv_inverse_iff [asymp_equiv_simps]: "(\x. inverse (f x)) \[F] (\x. inverse (g x)) \ f \[F] g" proof assume "(\x. inverse (f x)) \[F] (\x. inverse (g x))" hence "(\x. inverse (inverse (f x))) \[F] (\x. inverse (inverse (g x)))" (is ?P) by (rule asymp_equiv_inverse) also have "?P \ f \[F] g" by (intro asymp_equiv_cong) simp_all finally show "f \[F] g" . qed (simp_all add: asymp_equiv_inverse) lemma asymp_equiv_divide [asymp_equiv_intros]: assumes "f1 \[F] g1" "f2 \[F] g2" shows "(\x. f1 x / f2 x) \[F] (\x. g1 x / g2 x)" using asymp_equiv_mult[OF assms(1) asymp_equiv_inverse[OF assms(2)]] by (simp add: field_simps) lemma asymp_equiv_compose [asymp_equiv_intros]: assumes "f \[G] g" "filterlim h G F" shows "f \ h \[F] g \ h" proof - let ?T = "\f g x. if f x = 0 \ g x = 0 then 1 else f x / g x" have "f \ h \[F] g \ h \ ((?T f g \ h) \ 1) F" by (simp add: asymp_equiv_def o_def) also have "\ \ (?T f g \ 1) (filtermap h F)" by (rule tendsto_compose_filtermap) also have "\" by (rule tendsto_mono[of _ G]) (insert assms, simp_all add: asymp_equiv_def filterlim_def) finally show ?thesis . qed lemma asymp_equiv_compose': assumes "f \[G] g" "filterlim h G F" shows "(\x. f (h x)) \[F] (\x. g (h x))" using asymp_equiv_compose[OF assms] by (simp add: o_def) - + lemma asymp_equiv_powr_real [asymp_equiv_intros]: fixes f g :: "'a \ real" assumes "f \[F] g" "eventually (\x. f x \ 0) F" "eventually (\x. g x \ 0) F" shows "(\x. f x powr y) \[F] (\x. g x powr y)" proof - let ?T = "\f g x. if f x = 0 \ g x = 0 then 1 else f x / g x" have "((\x. ?T f g x powr y) \ 1 powr y) F" by (intro tendsto_intros asymp_equivD[OF assms(1)]) simp_all hence "((\x. ?T f g x powr y) \ 1) F" by simp moreover have "eventually (\x. ?T f g x powr y = ?T (\x. f x powr y) (\x. g x powr y) x) F" using asymp_equiv_eventually_zeros[OF assms(1)] assms(2,3) by eventually_elim (auto simp: powr_divide) ultimately show ?thesis unfolding asymp_equiv_def by (rule Lim_transform_eventually) qed lemma asymp_equiv_norm [asymp_equiv_intros]: fixes f g :: "'a \ 'b :: real_normed_field" assumes "f \[F] g" shows "(\x. norm (f x)) \[F] (\x. norm (g x))" - using tendsto_norm[OF asymp_equivD[OF assms]] + using tendsto_norm[OF asymp_equivD[OF assms]] by (simp add: if_distrib asymp_equiv_def norm_divide cong: if_cong) lemma asymp_equiv_abs_real [asymp_equiv_intros]: fixes f g :: "'a \ real" assumes "f \[F] g" shows "(\x. \f x\) \[F] (\x. \g x\)" - using tendsto_rabs[OF asymp_equivD[OF assms]] + using tendsto_rabs[OF asymp_equivD[OF assms]] by (simp add: if_distrib asymp_equiv_def cong: if_cong) lemma asymp_equiv_imp_eventually_le: assumes "f \[F] g" "c > 1" shows "eventually (\x. norm (f x) \ c * norm (g x)) F" proof - from order_tendstoD(2)[OF asymp_equivD[OF asymp_equiv_norm[OF assms(1)]] assms(2)] asymp_equiv_eventually_zeros[OF assms(1)] show ?thesis by eventually_elim (auto split: if_splits simp: field_simps) qed lemma asymp_equiv_imp_eventually_ge: assumes "f \[F] g" "c < 1" shows "eventually (\x. norm (f x) \ c * norm (g x)) F" proof - from order_tendstoD(1)[OF asymp_equivD[OF asymp_equiv_norm[OF assms(1)]] assms(2)] asymp_equiv_eventually_zeros[OF assms(1)] show ?thesis by eventually_elim (auto split: if_splits simp: field_simps) qed lemma asymp_equiv_imp_bigo: assumes "f \[F] g" shows "f \ O[F](g)" proof (rule bigoI) have "(3/2::real) > 1" by simp from asymp_equiv_imp_eventually_le[OF assms this] show "eventually (\x. norm (f x) \ 3/2 * norm (g x)) F" by eventually_elim simp qed lemma asymp_equiv_imp_bigomega: "f \[F] g \ f \ \[F](g)" using asymp_equiv_imp_bigo[of g F f] by (simp add: asymp_equiv_sym bigomega_iff_bigo) lemma asymp_equiv_imp_bigtheta: "f \[F] g \ f \ \[F](g)" by (intro bigthetaI asymp_equiv_imp_bigo asymp_equiv_imp_bigomega) lemma asymp_equiv_at_infinity_transfer: assumes "f \[F] g" "filterlim f at_infinity F" shows "filterlim g at_infinity F" proof - from assms(1) have "g \ \[F](f)" by (rule asymp_equiv_imp_bigtheta[OF asymp_equiv_symI]) also from assms have "f \ \[F](\_. 1)" by (simp add: smallomega_1_conv_filterlim) finally show ?thesis by (simp add: smallomega_1_conv_filterlim) qed lemma asymp_equiv_at_top_transfer: fixes f g :: "_ \ real" assumes "f \[F] g" "filterlim f at_top F" shows "filterlim g at_top F" proof (rule filterlim_at_infinity_imp_filterlim_at_top) show "filterlim g at_infinity F" by (rule asymp_equiv_at_infinity_transfer[OF assms(1) filterlim_mono[OF assms(2)]]) (auto simp: at_top_le_at_infinity) from assms(2) have "eventually (\x. f x > 0) F" using filterlim_at_top_dense by blast with asymp_equiv_eventually_pos_iff[OF assms(1)] show "eventually (\x. g x > 0) F" by eventually_elim blast qed lemma asymp_equiv_at_bot_transfer: fixes f g :: "_ \ real" assumes "f \[F] g" "filterlim f at_bot F" shows "filterlim g at_bot F" unfolding filterlim_uminus_at_bot by (rule asymp_equiv_at_top_transfer[of "\x. -f x" F "\x. -g x"]) - (insert assms, auto simp: filterlim_uminus_at_bot asymp_equiv_uminus) + (insert assms, auto simp: filterlim_uminus_at_bot asymp_equiv_uminus) lemma asymp_equivI'_const: assumes "((\x. f x / g x) \ c) F" "c \ 0" shows "f \[F] (\x. c * g x)" using tendsto_mult[OF assms(1) tendsto_const[of "inverse c"]] assms(2) by (intro asymp_equivI') (simp add: field_simps) lemma asymp_equivI'_inverse_const: assumes "((\x. f x / g x) \ inverse c) F" "c \ 0" shows "(\x. c * f x) \[F] g" using tendsto_mult[OF assms(1) tendsto_const[of "c"]] assms(2) by (intro asymp_equivI') (simp add: field_simps) lemma filterlim_at_bot_imp_at_infinity: "filterlim f at_bot F \ filterlim f at_infinity F" for f :: "_ \ real" using at_bot_le_at_infinity filterlim_mono by blast lemma asymp_equiv_imp_diff_smallo: assumes "f \[F] g" shows "(\x. f x - g x) \ o[F](g)" proof (rule landau_o.smallI) fix c :: real assume "c > 0" hence c: "min c 1 > 0" by simp let ?h = "\x. if f x = 0 \ g x = 0 then 1 else f x / g x" from assms have "((\x. ?h x - 1) \ 1 - 1) F" by (intro tendsto_diff asymp_equivD tendsto_const) from tendstoD[OF this c] show "eventually (\x. norm (f x - g x) \ c * norm (g x)) F" proof eventually_elim case (elim x) from elim have "norm (f x - g x) \ norm (f x / g x - 1) * norm (g x)" by (subst norm_mult [symmetric]) (auto split: if_splits simp add: algebra_simps) also have "norm (f x / g x - 1) * norm (g x) \ c * norm (g x)" using elim by (auto split: if_splits simp: mult_right_mono) finally show ?case . qed qed lemma asymp_equiv_altdef: "f \[F] g \ (\x. f x - g x) \ o[F](g)" by (rule iffI[OF asymp_equiv_imp_diff_smallo smallo_imp_asymp_equiv]) lemma asymp_equiv_0_left_iff [simp]: "(\_. 0) \[F] f \ eventually (\x. f x = 0) F" and asymp_equiv_0_right_iff [simp]: "f \[F] (\_. 0) \ eventually (\x. f x = 0) F" by (simp_all add: asymp_equiv_altdef landau_o.small_refl_iff) lemma asymp_equiv_sandwich_real: fixes f g l u :: "'a \ real" assumes "l \[F] g" "u \[F] g" "eventually (\x. f x \ {l x..u x}) F" shows "f \[F] g" unfolding asymp_equiv_altdef proof (rule landau_o.smallI) fix c :: real assume c: "c > 0" have "eventually (\x. norm (f x - g x) \ max (norm (l x - g x)) (norm (u x - g x))) F" using assms(3) by eventually_elim auto moreover have "eventually (\x. norm (l x - g x) \ c * norm (g x)) F" "eventually (\x. norm (u x - g x) \ c * norm (g x)) F" using assms(1,2) by (auto simp: asymp_equiv_altdef dest: landau_o.smallD[OF _ c]) hence "eventually (\x. max (norm (l x - g x)) (norm (u x - g x)) \ c * norm (g x)) F" by eventually_elim simp ultimately show "eventually (\x. norm (f x - g x) \ c * norm (g x)) F" by eventually_elim (rule order.trans) qed lemma asymp_equiv_sandwich_real': fixes f g l u :: "'a \ real" assumes "f \[F] l" "f \[F] u" "eventually (\x. g x \ {l x..u x}) F" shows "f \[F] g" using asymp_equiv_sandwich_real[of l F f u g] assms by (simp add: asymp_equiv_sym) lemma asymp_equiv_sandwich_real'': fixes f g l u :: "'a \ real" assumes "l1 \[F] u1" "u1 \[F] l2" "l2 \[F] u2" "eventually (\x. f x \ {l1 x..u1 x}) F" "eventually (\x. g x \ {l2 x..u2 x}) F" shows "f \[F] g" by (rule asymp_equiv_sandwich_real[OF asymp_equiv_sandwich_real'[OF _ _ assms(5)] asymp_equiv_sandwich_real'[OF _ _ assms(5)] assms(4)]; blast intro: asymp_equiv_trans assms(1,2,3))+ end diff --git a/src/HOL/Tools/Argo/argo_real.ML b/src/HOL/Tools/Argo/argo_real.ML --- a/src/HOL/Tools/Argo/argo_real.ML +++ b/src/HOL/Tools/Argo/argo_real.ML @@ -1,333 +1,332 @@ (* Title: HOL/Tools/Argo/argo_real.ML Author: Sascha Boehme Extension of the Argo tactic for the reals. *) structure Argo_Real: sig end = struct (* translating input terms *) -fun trans_type _ \<^typ>\Real.real\ tcx = SOME (Argo_Expr.Real, tcx) +fun trans_type _ \<^Type>\real\ tcx = SOME (Argo_Expr.Real, tcx) | trans_type _ _ _ = NONE -fun trans_term f (@{const Groups.uminus_class.uminus (real)} $ t) tcx = +fun trans_term f \<^Const_>\uminus \\<^Type>\real\\ for t\ tcx = tcx |> f t |>> Argo_Expr.mk_neg |> SOME - | trans_term f (@{const Groups.plus_class.plus (real)} $ t1 $ t2) tcx = + | trans_term f \<^Const_>\plus \\<^Type>\real\\ for t1 t2\ tcx = tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_add2 |> SOME - | trans_term f (@{const Groups.minus_class.minus (real)} $ t1 $ t2) tcx = + | trans_term f \<^Const_>\minus \\<^Type>\real\\ for t1 t2\ tcx = tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_sub |> SOME - | trans_term f (@{const Groups.times_class.times (real)} $ t1 $ t2) tcx = + | trans_term f \<^Const_>\times \\<^Type>\real\\ for t1 t2\ tcx = tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_mul |> SOME - | trans_term f (@{const Rings.divide_class.divide (real)} $ t1 $ t2) tcx = + | trans_term f \<^Const_>\divide \\<^Type>\real\\ for t1 t2\ tcx = tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_div |> SOME - | trans_term f (@{const Orderings.ord_class.min (real)} $ t1 $ t2) tcx = + | trans_term f \<^Const_>\min \\<^Type>\real\\ for t1 t2\ tcx = tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_min |> SOME - | trans_term f (@{const Orderings.ord_class.max (real)} $ t1 $ t2) tcx = + | trans_term f \<^Const_>\max \\<^Type>\real\\ for t1 t2\ tcx = tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_max |> SOME - | trans_term f (@{const Groups.abs_class.abs (real)} $ t) tcx = + | trans_term f \<^Const_>\abs \\<^Type>\real\\ for t\ tcx = tcx |> f t |>> Argo_Expr.mk_abs |> SOME - | trans_term f (@{const Orderings.ord_class.less (real)} $ t1 $ t2) tcx = + | trans_term f \<^Const_>\less \\<^Type>\real\\ for t1 t2\ tcx = tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_lt |> SOME - | trans_term f (@{const Orderings.ord_class.less_eq (real)} $ t1 $ t2) tcx = + | trans_term f \<^Const_>\less_eq \\<^Type>\real\\ for t1 t2\ tcx = tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_le |> SOME | trans_term _ t tcx = (case try HOLogic.dest_number t of - SOME (\<^typ>\Real.real\, n) => SOME (Argo_Expr.mk_num (Rat.of_int n), tcx) + SOME (\<^Type>\real\, n) => SOME (Argo_Expr.mk_num (Rat.of_int n), tcx) | _ => NONE) (* reverse translation *) -fun mk_plus t1 t2 = @{const Groups.plus_class.plus (real)} $ t1 $ t2 +fun mk_plus t1 t2 = \<^Const>\plus \\<^Type>\real\\ for t1 t2\ fun mk_sum ts = uncurry (fold_rev mk_plus) (split_last ts) -fun mk_times t1 t2 = @{const Groups.times_class.times (real)} $ t1 $ t2 -fun mk_divide t1 t2 = @{const Rings.divide_class.divide (real)} $ t1 $ t2 -fun mk_le t1 t2 = @{const Orderings.ord_class.less_eq (real)} $ t1 $ t2 -fun mk_lt t1 t2 = @{const Orderings.ord_class.less (real)} $ t1 $ t2 +fun mk_times t1 t2 = \<^Const>\times \\<^Type>\real\\ for t1 t2\ +fun mk_divide t1 t2 = \<^Const>\divide \\<^Type>\real\\ for t1 t2\ +fun mk_le t1 t2 = \<^Const>\less_eq \\<^Type>\real\\ for t1 t2\ +fun mk_lt t1 t2 = \<^Const>\less \\<^Type>\real\\ for t1 t2\ -fun mk_real_num i = HOLogic.mk_number \<^typ>\Real.real\ i +fun mk_real_num i = HOLogic.mk_number \<^Type>\real\ i fun mk_number n = let val (p, q) = Rat.dest n in if q = 1 then mk_real_num p else mk_divide (mk_real_num p) (mk_real_num q) end fun term_of _ (Argo_Expr.E (Argo_Expr.Num n, _)) = SOME (mk_number n) - | term_of f (Argo_Expr.E (Argo_Expr.Neg, [e])) = - SOME (@{const Groups.uminus_class.uminus (real)} $ f e) + | term_of f (Argo_Expr.E (Argo_Expr.Neg, [e])) = SOME \<^Const>\uminus \\<^Type>\real\\ for \f e\\ | term_of f (Argo_Expr.E (Argo_Expr.Add, es)) = SOME (mk_sum (map f es)) | term_of f (Argo_Expr.E (Argo_Expr.Sub, [e1, e2])) = - SOME (@{const Groups.minus_class.minus (real)} $ f e1 $ f e2) + SOME \<^Const>\minus \\<^Type>\real\\ for \f e1\ \f e2\\ | term_of f (Argo_Expr.E (Argo_Expr.Mul, [e1, e2])) = SOME (mk_times (f e1) (f e2)) | term_of f (Argo_Expr.E (Argo_Expr.Div, [e1, e2])) = SOME (mk_divide (f e1) (f e2)) | term_of f (Argo_Expr.E (Argo_Expr.Min, [e1, e2])) = - SOME (@{const Orderings.ord_class.min (real)} $ f e1 $ f e2) + SOME \<^Const>\min \\<^Type>\real\\ for \f e1\ \f e2\\ | term_of f (Argo_Expr.E (Argo_Expr.Max, [e1, e2])) = - SOME (@{const Orderings.ord_class.max (real)} $ f e1 $ f e2) - | term_of f (Argo_Expr.E (Argo_Expr.Abs, [e])) = SOME (@{const Groups.abs_class.abs (real)} $ f e) + SOME \<^Const>\max \\<^Type>\real\\ for \f e1\ \f e2\\ + | term_of f (Argo_Expr.E (Argo_Expr.Abs, [e])) = SOME \<^Const>\abs \\<^Type>\real\\ for \f e\\ | term_of f (Argo_Expr.E (Argo_Expr.Le, [e1, e2])) = SOME (mk_le (f e1) (f e2)) | term_of f (Argo_Expr.E (Argo_Expr.Lt, [e1, e2])) = SOME (mk_lt (f e1) (f e2)) | term_of _ _ = NONE (* proof replay for rewrite steps *) fun mk_rewr thm = thm RS @{thm eq_reflection} fun by_simp ctxt t = let fun prove {context, ...} = HEADGOAL (Simplifier.simp_tac context) in Goal.prove ctxt [] [] (HOLogic.mk_Trueprop t) prove end fun prove_num_pred ctxt n = by_simp ctxt (uncurry mk_lt (apply2 mk_number (if @0 < n then (@0, n) else (n, @0)))) fun simp_conv ctxt t = Conv.rewr_conv (mk_rewr (by_simp ctxt t)) fun nums_conv mk f ctxt n m = simp_conv ctxt (HOLogic.mk_eq (mk (mk_number n) (mk_number m), mk_number (f (n, m)))) val add_nums_conv = nums_conv mk_plus (op +) val mul_nums_conv = nums_conv mk_times (op *) val div_nums_conv = nums_conv mk_divide (op /) fun inv_num_conv ctxt = nums_conv mk_divide (fn (_, n) => Rat.inv n) ctxt @1 fun cmp_nums_conv ctxt b ct = - let val t = if b then \<^const>\HOL.True\ else \<^const>\HOL.False\ + let val t = if b then \<^Const>\True\ else \<^Const>\False\ in simp_conv ctxt (HOLogic.mk_eq (Thm.term_of ct, t)) ct end local -fun is_add2 (@{const Groups.plus_class.plus (real)} $ _ $ _) = true +fun is_add2 \<^Const_>\plus \\<^Type>\real\\ for _ _\ = true | is_add2 _ = false -fun is_add3 (@{const Groups.plus_class.plus (real)} $ _ $ t) = is_add2 t +fun is_add3 \<^Const_>\plus \\<^Type>\real\\ for _ t\ = is_add2 t | is_add3 _ = false val flatten_thm = mk_rewr @{lemma "(a::real) + b + c = a + (b + c)" by simp} fun flatten_conv ct = if is_add2 (Thm.term_of ct) then Argo_Tactic.flatten_conv flatten_conv flatten_thm ct else Conv.all_conv ct val swap_conv = Conv.rewrs_conv (map mk_rewr @{lemma "(a::real) + (b + c) = b + (a + c)" "(a::real) + b = b + a" by simp_all}) val assoc_conv = Conv.rewr_conv (mk_rewr @{lemma "(a::real) + (b + c) = (a + b) + c" by simp}) val norm_monom_thm = mk_rewr @{lemma "1 * (a::real) = a" by simp} fun norm_monom_conv n = if n = @1 then Conv.rewr_conv norm_monom_thm else Conv.all_conv val add2_thms = map mk_rewr @{lemma "n * (a::real) + m * a = (n + m) * a" "n * (a::real) + a = (n + 1) * a" "(a::real) + m * a = (1 + m) * a" "(a::real) + a = (1 + 1) * a" by algebra+} val add3_thms = map mk_rewr @{lemma "n * (a::real) + (m * a + b) = (n + m) * a + b" "n * (a::real) + (a + b) = (n + 1) * a + b" "(a::real) + (m * a + b) = (1 + m) * a + b" "(a::real) + (a + b) = (1 + 1) * a + b" by algebra+} fun choose_conv cv2 cv3 ct = if is_add3 (Thm.term_of ct) then cv3 ct else cv2 ct fun join_num_conv ctxt n m = let val conv = add_nums_conv ctxt n m in choose_conv conv (assoc_conv then_conv Conv.arg1_conv conv) end fun join_monom_conv ctxt n m = let val conv = Conv.arg1_conv (add_nums_conv ctxt n m) then_conv norm_monom_conv (n + m) fun seq_conv thms cv = Conv.rewrs_conv thms then_conv cv in choose_conv (seq_conv add2_thms conv) (seq_conv add3_thms (Conv.arg1_conv conv)) end fun join_conv NONE = join_num_conv | join_conv (SOME _) = join_monom_conv fun bubble_down_conv _ _ [] ct = Conv.no_conv ct | bubble_down_conv _ _ [_] ct = Conv.all_conv ct | bubble_down_conv ctxt i ((m1 as (n1, i1)) :: (m2 as (n2, i2)) :: ms) ct = if i1 = i then if i2 = i then (join_conv i ctxt n1 n2 then_conv bubble_down_conv ctxt i ((n1 + n2, i) :: ms)) ct else (swap_conv then_conv Conv.arg_conv (bubble_down_conv ctxt i (m1 :: ms))) ct else Conv.arg_conv (bubble_down_conv ctxt i (m2 :: ms)) ct fun drop_var i ms = filter_out (fn (_, i') => i' = i) ms fun permute_conv _ [] [] ct = Conv.all_conv ct | permute_conv ctxt (ms as ((_, i) :: _)) [] ct = (bubble_down_conv ctxt i ms then_conv permute_conv ctxt (drop_var i ms) []) ct | permute_conv ctxt ms1 ms2 ct = let val (ms2', (_, i)) = split_last ms2 in (bubble_down_conv ctxt i ms1 then_conv permute_conv ctxt (drop_var i ms1) ms2') ct end val no_monom_conv = Conv.rewr_conv (mk_rewr @{lemma "0 * (a::real) = 0" by simp}) val norm_sum_conv = Conv.rewrs_conv (map mk_rewr @{lemma "0 * (a::real) + b = b" "(a::real) + 0 * b = a" "0 + (a::real) = a" "(a::real) + 0 = a" by simp_all}) fun drop0_conv ct = if is_add2 (Thm.term_of ct) then ((norm_sum_conv then_conv drop0_conv) else_conv Conv.arg_conv drop0_conv) ct else Conv.try_conv no_monom_conv ct fun full_add_conv ctxt ms1 ms2 = if eq_list (op =) (ms1, ms2) then flatten_conv else flatten_conv then_conv permute_conv ctxt ms1 ms2 then_conv drop0_conv in fun add_conv ctxt (ms1, ms2 as [(n, NONE)]) = if n = @0 then full_add_conv ctxt ms1 [] else full_add_conv ctxt ms1 ms2 | add_conv ctxt (ms1, ms2) = full_add_conv ctxt ms1 ms2 end val mul_suml_thm = mk_rewr @{lemma "((x::real) + y) * z = x * z + y * z" by (rule ring_distribs)} val mul_sumr_thm = mk_rewr @{lemma "(x::real) * (y + z) = x * y + x * z" by (rule ring_distribs)} fun mul_sum_conv thm ct = Conv.try_conv (Conv.rewr_conv thm then_conv Conv.binop_conv (mul_sum_conv thm)) ct val div_sum_thm = mk_rewr @{lemma "((x::real) + y) / z = x / z + y / z" by (rule add_divide_distrib)} fun div_sum_conv ct = Conv.try_conv (Conv.rewr_conv div_sum_thm then_conv Conv.binop_conv div_sum_conv) ct fun var_of_add_cmp (_ $ _ $ (_ $ _ $ (_ $ _ $ Var v))) = v | var_of_add_cmp t = raise TERM ("var_of_add_cmp", [t]) fun add_cmp_conv ctxt n thm = let val v = var_of_add_cmp (Thm.prop_of thm) in Conv.rewr_conv (Thm.instantiate (TVars.empty, Vars.make [(v, Thm.cterm_of ctxt (mk_number n))]) thm) end fun mul_cmp_conv ctxt n pos_thm neg_thm = let val thm = if @0 < n then pos_thm else neg_thm in Conv.rewr_conv (prove_num_pred ctxt n RS thm) end val neg_thm = mk_rewr @{lemma "-(x::real) = -1 * x" by simp} val sub_thm = mk_rewr @{lemma "(x::real) - y = x + -1 * y" by simp} val mul_zero_thm = mk_rewr @{lemma "0 * (x::real) = 0" by (rule mult_zero_left)} val mul_one_thm = mk_rewr @{lemma "1 * (x::real) = x" by (rule mult_1)} val mul_comm_thm = mk_rewr @{lemma "(x::real) * y = y * x" by simp} val mul_assocl_thm = mk_rewr @{lemma "((x::real) * y) * z = x * (y * z)" by simp} val mul_assocr_thm = mk_rewr @{lemma "(x::real) * (y * z) = (x * y) * z" by simp} val mul_divl_thm = mk_rewr @{lemma "((x::real) / y) * z = (x * z) / y" by simp} val mul_divr_thm = mk_rewr @{lemma "(x::real) * (y / z) = (x * y) / z" by simp} val div_zero_thm = mk_rewr @{lemma "0 / (x::real) = 0" by (rule div_0)} val div_one_thm = mk_rewr @{lemma "(x::real) / 1 = x" by (rule div_by_1)} val div_numl_thm = mk_rewr @{lemma "(x::real) / y = x * (1 / y)" by simp} val div_numr_thm = mk_rewr @{lemma "(x::real) / y = (1 / y) * x" by simp} val div_mull_thm = mk_rewr @{lemma "((x::real) * y) / z = x * (y / z)" by simp} val div_mulr_thm = mk_rewr @{lemma "(x::real) / (y * z) = (1 / y) * (x / z)" by simp} val div_divl_thm = mk_rewr @{lemma "((x::real) / y) / z = x / (y * z)" by simp} val div_divr_thm = mk_rewr @{lemma "(x::real) / (y / z) = (x * z) / y" by simp} val min_eq_thm = mk_rewr @{lemma "min (x::real) x = x" by (simp add: min_def)} val min_lt_thm = mk_rewr @{lemma "min (x::real) y = (if x <= y then x else y)" by (rule min_def)} val min_gt_thm = mk_rewr @{lemma "min (x::real) y = (if y <= x then y else x)" by (simp add: min_def)} val max_eq_thm = mk_rewr @{lemma "max (x::real) x = x" by (simp add: max_def)} val max_lt_thm = mk_rewr @{lemma "max (x::real) y = (if x <= y then y else x)" by (rule max_def)} val max_gt_thm = mk_rewr @{lemma "max (x::real) y = (if y <= x then x else y)" by (simp add: max_def)} val abs_thm = mk_rewr @{lemma "abs (x::real) = (if 0 <= x then x else -x)" by simp} val sub_eq_thm = mk_rewr @{lemma "((x::real) = y) = (x - y = 0)" by simp} val eq_le_thm = mk_rewr @{lemma "((x::real) = y) = ((x \ y) \ (y \ x))" by (rule order_eq_iff)} val add_le_thm = mk_rewr @{lemma "((x::real) \ y) = (x + n \ y + n)" by simp} val add_lt_thm = mk_rewr @{lemma "((x::real) < y) = (x + n < y + n)" by simp} val sub_le_thm = mk_rewr @{lemma "((x::real) <= y) = (x - y <= 0)" by simp} val sub_lt_thm = mk_rewr @{lemma "((x::real) < y) = (x - y < 0)" by simp} val pos_mul_le_thm = mk_rewr @{lemma "0 < n ==> ((x::real) <= y) = (n * x <= n * y)" by simp} val neg_mul_le_thm = mk_rewr @{lemma "n < 0 ==> ((x::real) <= y) = (n * y <= n * x)" by simp} val pos_mul_lt_thm = mk_rewr @{lemma "0 < n ==> ((x::real) < y) = (n * x < n * y)" by simp} val neg_mul_lt_thm = mk_rewr @{lemma "n < 0 ==> ((x::real) < y) = (n * y < n * x)" by simp} val not_le_thm = mk_rewr @{lemma "(\((x::real) \ y)) = (y < x)" by auto} val not_lt_thm = mk_rewr @{lemma "(\((x::real) < y)) = (y \ x)" by auto} fun replay_rewr _ Argo_Proof.Rewr_Neg = Conv.rewr_conv neg_thm | replay_rewr ctxt (Argo_Proof.Rewr_Add ps) = add_conv ctxt ps | replay_rewr _ Argo_Proof.Rewr_Sub = Conv.rewr_conv sub_thm | replay_rewr _ Argo_Proof.Rewr_Mul_Zero = Conv.rewr_conv mul_zero_thm | replay_rewr _ Argo_Proof.Rewr_Mul_One = Conv.rewr_conv mul_one_thm | replay_rewr ctxt (Argo_Proof.Rewr_Mul_Nums (n, m)) = mul_nums_conv ctxt n m | replay_rewr _ Argo_Proof.Rewr_Mul_Comm = Conv.rewr_conv mul_comm_thm | replay_rewr _ (Argo_Proof.Rewr_Mul_Assoc Argo_Proof.Left) = Conv.rewr_conv mul_assocl_thm | replay_rewr _ (Argo_Proof.Rewr_Mul_Assoc Argo_Proof.Right) = Conv.rewr_conv mul_assocr_thm | replay_rewr _ (Argo_Proof.Rewr_Mul_Sum Argo_Proof.Left) = mul_sum_conv mul_suml_thm | replay_rewr _ (Argo_Proof.Rewr_Mul_Sum Argo_Proof.Right) = mul_sum_conv mul_sumr_thm | replay_rewr _ (Argo_Proof.Rewr_Mul_Div Argo_Proof.Left) = Conv.rewr_conv mul_divl_thm | replay_rewr _ (Argo_Proof.Rewr_Mul_Div Argo_Proof.Right) = Conv.rewr_conv mul_divr_thm | replay_rewr _ Argo_Proof.Rewr_Div_Zero = Conv.rewr_conv div_zero_thm | replay_rewr _ Argo_Proof.Rewr_Div_One = Conv.rewr_conv div_one_thm | replay_rewr ctxt (Argo_Proof.Rewr_Div_Nums (n, m)) = div_nums_conv ctxt n m | replay_rewr _ (Argo_Proof.Rewr_Div_Num (Argo_Proof.Left, _)) = Conv.rewr_conv div_numl_thm | replay_rewr ctxt (Argo_Proof.Rewr_Div_Num (Argo_Proof.Right, n)) = Conv.rewr_conv div_numr_thm then_conv Conv.arg1_conv (inv_num_conv ctxt n) | replay_rewr _ (Argo_Proof.Rewr_Div_Mul (Argo_Proof.Left, _)) = Conv.rewr_conv div_mull_thm | replay_rewr ctxt (Argo_Proof.Rewr_Div_Mul (Argo_Proof.Right, n)) = Conv.rewr_conv div_mulr_thm then_conv Conv.arg1_conv (inv_num_conv ctxt n) | replay_rewr _ (Argo_Proof.Rewr_Div_Div Argo_Proof.Left) = Conv.rewr_conv div_divl_thm | replay_rewr _ (Argo_Proof.Rewr_Div_Div Argo_Proof.Right) = Conv.rewr_conv div_divr_thm | replay_rewr _ Argo_Proof.Rewr_Div_Sum = div_sum_conv | replay_rewr _ Argo_Proof.Rewr_Min_Eq = Conv.rewr_conv min_eq_thm | replay_rewr _ Argo_Proof.Rewr_Min_Lt = Conv.rewr_conv min_lt_thm | replay_rewr _ Argo_Proof.Rewr_Min_Gt = Conv.rewr_conv min_gt_thm | replay_rewr _ Argo_Proof.Rewr_Max_Eq = Conv.rewr_conv max_eq_thm | replay_rewr _ Argo_Proof.Rewr_Max_Lt = Conv.rewr_conv max_lt_thm | replay_rewr _ Argo_Proof.Rewr_Max_Gt = Conv.rewr_conv max_gt_thm | replay_rewr _ Argo_Proof.Rewr_Abs = Conv.rewr_conv abs_thm | replay_rewr ctxt (Argo_Proof.Rewr_Eq_Nums b) = cmp_nums_conv ctxt b | replay_rewr _ Argo_Proof.Rewr_Eq_Sub = Conv.rewr_conv sub_eq_thm | replay_rewr _ Argo_Proof.Rewr_Eq_Le = Conv.rewr_conv eq_le_thm | replay_rewr ctxt (Argo_Proof.Rewr_Ineq_Nums (_, b)) = cmp_nums_conv ctxt b | replay_rewr ctxt (Argo_Proof.Rewr_Ineq_Add (Argo_Proof.Le, n)) = add_cmp_conv ctxt n add_le_thm | replay_rewr ctxt (Argo_Proof.Rewr_Ineq_Add (Argo_Proof.Lt, n)) = add_cmp_conv ctxt n add_lt_thm | replay_rewr _ (Argo_Proof.Rewr_Ineq_Sub Argo_Proof.Le) = Conv.rewr_conv sub_le_thm | replay_rewr _ (Argo_Proof.Rewr_Ineq_Sub Argo_Proof.Lt) = Conv.rewr_conv sub_lt_thm | replay_rewr ctxt (Argo_Proof.Rewr_Ineq_Mul (Argo_Proof.Le, n)) = mul_cmp_conv ctxt n pos_mul_le_thm neg_mul_le_thm | replay_rewr ctxt (Argo_Proof.Rewr_Ineq_Mul (Argo_Proof.Lt, n)) = mul_cmp_conv ctxt n pos_mul_lt_thm neg_mul_lt_thm | replay_rewr _ (Argo_Proof.Rewr_Not_Ineq Argo_Proof.Le) = Conv.rewr_conv not_le_thm | replay_rewr _ (Argo_Proof.Rewr_Not_Ineq Argo_Proof.Lt) = Conv.rewr_conv not_lt_thm | replay_rewr _ _ = Conv.no_conv (* proof replay *) val combine_thms = @{lemma "(a::real) < b ==> c < d ==> a + c < b + d" "(a::real) < b ==> c <= d ==> a + c < b + d" "(a::real) <= b ==> c < d ==> a + c < b + d" "(a::real) <= b ==> c <= d ==> a + c <= b + d" by auto} fun combine thm1 thm2 = hd (Argo_Tactic.discharges thm2 (Argo_Tactic.discharges thm1 combine_thms)) fun replay _ _ Argo_Proof.Linear_Comb prems = SOME (uncurry (fold_rev combine) (split_last prems)) | replay _ _ _ _ = NONE (* real extension of the Argo solver *) val _ = Theory.setup (Argo_Tactic.add_extension { trans_type = trans_type, trans_term = trans_term, term_of = term_of, replay_rewr = replay_rewr, replay = replay}) end diff --git a/src/HOL/Tools/Argo/argo_tactic.ML b/src/HOL/Tools/Argo/argo_tactic.ML --- a/src/HOL/Tools/Argo/argo_tactic.ML +++ b/src/HOL/Tools/Argo/argo_tactic.ML @@ -1,733 +1,731 @@ (* Title: HOL/Tools/Argo/argo_tactic.ML Author: Sascha Boehme HOL method and tactic for the Argo solver. *) signature ARGO_TACTIC = sig val trace: string Config.T val timeout: real Config.T (* extending the tactic *) type trans_context = Name.context * Argo_Expr.typ Typtab.table * (string * Argo_Expr.typ) Termtab.table type ('a, 'b) trans = 'a -> trans_context -> 'b * trans_context type ('a, 'b) trans' = 'a -> trans_context -> ('b * trans_context) option type extension = { trans_type: (typ, Argo_Expr.typ) trans -> (typ, Argo_Expr.typ) trans', trans_term: (term, Argo_Expr.expr) trans -> (term, Argo_Expr.expr) trans', term_of: (Argo_Expr.expr -> term) -> Argo_Expr.expr -> term option, replay_rewr: Proof.context -> Argo_Proof.rewrite -> conv, replay: (Argo_Expr.expr -> cterm) -> Proof.context -> Argo_Proof.rule -> thm list -> thm option} val add_extension: extension -> theory -> theory (* proof utilities *) val discharges: thm -> thm list -> thm list val flatten_conv: conv -> thm -> conv (* interface to the tactic as well as the underlying checker and prover *) datatype result = Satisfiable of term -> bool option | Unsatisfiable val check: term list -> Proof.context -> result * Proof.context val prove: thm list -> Proof.context -> thm option * Proof.context val argo_tac: Proof.context -> thm list -> int -> tactic end structure Argo_Tactic: ARGO_TACTIC = struct (* readable fresh names for terms *) fun fresh_name n = Name.variant (case Long_Name.base_name n of "" => "x" | n' => n') fun fresh_type_name (Type (n, _)) = fresh_name n | fresh_type_name (TFree (n, _)) = fresh_name n | fresh_type_name (TVar ((n, i), _)) = fresh_name (n ^ "." ^ string_of_int i) fun fresh_term_name (Const (n, _)) = fresh_name n | fresh_term_name (Free (n, _)) = fresh_name n | fresh_term_name (Var ((n, i), _)) = fresh_name (n ^ "." ^ string_of_int i) | fresh_term_name _ = fresh_name "" (* tracing *) datatype mode = None | Basic | Full fun string_of_mode None = "none" | string_of_mode Basic = "basic" | string_of_mode Full = "full" fun requires_mode None = [] | requires_mode Basic = [Basic, Full] | requires_mode Full = [Full] val trace = Attrib.setup_config_string \<^binding>\argo_trace\ (K (string_of_mode None)) fun allows_mode ctxt = exists (equal (Config.get ctxt trace) o string_of_mode) o requires_mode fun output mode ctxt msg = if allows_mode ctxt mode then Output.tracing ("Argo: " ^ msg) else () val tracing = output Basic val full_tracing = output Full fun with_mode mode ctxt f = if allows_mode ctxt mode then f ctxt else () val with_tracing = with_mode Basic val with_full_tracing = with_mode Full (* timeout *) val timeout = Attrib.setup_config_real \<^binding>\argo_timeout\ (K 10.0) val timeout_seconds = seconds o Config.apply timeout fun with_timeout ctxt f x = Timeout.apply (timeout_seconds ctxt) f x (* extending the tactic *) type trans_context = Name.context * Argo_Expr.typ Typtab.table * (string * Argo_Expr.typ) Termtab.table type ('a, 'b) trans = 'a -> trans_context -> 'b * trans_context type ('a, 'b) trans' = 'a -> trans_context -> ('b * trans_context) option type extension = { trans_type: (typ, Argo_Expr.typ) trans -> (typ, Argo_Expr.typ) trans', trans_term: (term, Argo_Expr.expr) trans -> (term, Argo_Expr.expr) trans', term_of: (Argo_Expr.expr -> term) -> Argo_Expr.expr -> term option, replay_rewr: Proof.context -> Argo_Proof.rewrite -> conv, replay: (Argo_Expr.expr -> cterm) -> Proof.context -> Argo_Proof.rule -> thm list -> thm option} fun eq_extension ((serial1, _), (serial2, _)) = (serial1 = serial2) structure Extensions = Theory_Data ( type T = (serial * extension) list val empty = [] val extend = I val merge = merge eq_extension ) fun add_extension ext = Extensions.map (insert eq_extension (serial (), ext)) fun get_extensions ctxt = Extensions.get (Proof_Context.theory_of ctxt) fun apply_first ctxt f = get_first (fn (_, e) => f e) (get_extensions ctxt) fun ext_trans sel ctxt f x tcx = apply_first ctxt (fn ext => sel ext f x tcx) val ext_trans_type = ext_trans (fn {trans_type, ...}: extension => trans_type) val ext_trans_term = ext_trans (fn {trans_term, ...}: extension => trans_term) fun ext_term_of ctxt f e = apply_first ctxt (fn {term_of, ...}: extension => term_of f e) fun ext_replay_rewr ctxt r = get_extensions ctxt |> map (fn (_, {replay_rewr, ...}: extension) => replay_rewr ctxt r) |> Conv.first_conv fun ext_replay cprop_of ctxt rule prems = (case apply_first ctxt (fn {replay, ...}: extension => replay cprop_of ctxt rule prems) of SOME thm => thm | NONE => raise THM ("failed to replay " ^ quote (Argo_Proof.string_of_rule rule), 0, [])) (* translating input terms *) fun add_new_type T (names, types, terms) = let val (n, names') = fresh_type_name T names val ty = Argo_Expr.Type n in (ty, (names', Typtab.update (T, ty) types, terms)) end fun add_type T (tcx as (_, types, _)) = (case Typtab.lookup types T of SOME ty => (ty, tcx) | NONE => add_new_type T tcx) -fun trans_type _ \<^typ>\HOL.bool\ = pair Argo_Expr.Bool - | trans_type ctxt (Type (\<^type_name>\fun\, [T1, T2])) = +fun trans_type _ \<^Type>\bool\ = pair Argo_Expr.Bool + | trans_type ctxt \<^Type>\fun T1 T2\ = trans_type ctxt T1 ##>> trans_type ctxt T2 #>> Argo_Expr.Func | trans_type ctxt T = (fn tcx => (case ext_trans_type ctxt (trans_type ctxt) T tcx of SOME result => result | NONE => add_type T tcx)) fun add_new_term ctxt t T tcx = let val (ty, (names, types, terms)) = trans_type ctxt T tcx val (n, names') = fresh_term_name t names val c = (n, ty) in (Argo_Expr.mk_con c, (names', types, Termtab.update (t, c) terms)) end fun add_term ctxt t (tcx as (_, _, terms)) = (case Termtab.lookup terms t of SOME c => (Argo_Expr.mk_con c, tcx) | NONE => add_new_term ctxt t (Term.fastype_of t) tcx) -fun mk_eq \<^typ>\HOL.bool\ = Argo_Expr.mk_iff +fun mk_eq \<^Type>\bool\ = Argo_Expr.mk_iff | mk_eq _ = Argo_Expr.mk_eq -fun trans_term _ \<^const>\HOL.True\ = pair Argo_Expr.true_expr - | trans_term _ \<^const>\HOL.False\ = pair Argo_Expr.false_expr - | trans_term ctxt (\<^const>\HOL.Not\ $ t) = trans_term ctxt t #>> Argo_Expr.mk_not - | trans_term ctxt (\<^const>\HOL.conj\ $ t1 $ t2) = +fun trans_term _ \<^Const_>\True\ = pair Argo_Expr.true_expr + | trans_term _ \<^Const_>\False\ = pair Argo_Expr.false_expr + | trans_term ctxt \<^Const_>\Not for t\ = trans_term ctxt t #>> Argo_Expr.mk_not + | trans_term ctxt \<^Const_>\conj for t1 t2\ = trans_term ctxt t1 ##>> trans_term ctxt t2 #>> uncurry Argo_Expr.mk_and2 - | trans_term ctxt (\<^const>\HOL.disj\ $ t1 $ t2) = + | trans_term ctxt \<^Const_>\disj for t1 t2\ = trans_term ctxt t1 ##>> trans_term ctxt t2 #>> uncurry Argo_Expr.mk_or2 - | trans_term ctxt (\<^const>\HOL.implies\ $ t1 $ t2) = + | trans_term ctxt \<^Const_>\implies for t1 t2\ = trans_term ctxt t1 ##>> trans_term ctxt t2 #>> uncurry Argo_Expr.mk_imp - | trans_term ctxt (Const (\<^const_name>\HOL.If\, _) $ t1 $ t2 $ t3) = + | trans_term ctxt \<^Const_>\If _ for t1 t2 t3\ = trans_term ctxt t1 ##>> trans_term ctxt t2 ##>> trans_term ctxt t3 #>> (fn ((u1, u2), u3) => Argo_Expr.mk_ite u1 u2 u3) - | trans_term ctxt (Const (\<^const_name>\HOL.eq\, T) $ t1 $ t2) = - trans_term ctxt t1 ##>> trans_term ctxt t2 #>> uncurry (mk_eq (Term.domain_type T)) + | trans_term ctxt \<^Const_>\HOL.eq T for t1 t2\ = + trans_term ctxt t1 ##>> trans_term ctxt t2 #>> uncurry (mk_eq T) | trans_term ctxt (t as (t1 $ t2)) = (fn tcx => (case ext_trans_term ctxt (trans_term ctxt) t tcx of SOME result => result | NONE => tcx |> trans_term ctxt t1 ||>> trans_term ctxt t2 |>> uncurry Argo_Expr.mk_app)) | trans_term ctxt t = (fn tcx => (case ext_trans_term ctxt (trans_term ctxt) t tcx of SOME result => result | NONE => add_term ctxt t tcx)) fun translate ctxt prop = trans_term ctxt (HOLogic.dest_Trueprop prop) (* invoking the solver *) type data = { names: Name.context, types: Argo_Expr.typ Typtab.table, terms: (string * Argo_Expr.typ) Termtab.table, argo: Argo_Solver.context} fun mk_data names types terms argo: data = {names=names, types=types, terms=terms, argo=argo} val empty_data = mk_data Name.context Typtab.empty Termtab.empty Argo_Solver.context structure Solver_Data = Proof_Data ( type T = data option fun init _ = SOME empty_data ) datatype ('m, 'p) solver_result = Model of 'm | Proof of 'p fun raw_solve es argo = Model (Argo_Solver.assert es argo) handle Argo_Proof.UNSAT proof => Proof proof fun value_of terms model t = (case Termtab.lookup terms t of SOME c => model c | _ => NONE) fun trace_props props ctxt = tracing ctxt (Pretty.string_of (Pretty.big_list "using these propositions:" (map (Syntax.pretty_term ctxt) props))) fun trace_result ctxt ({elapsed, ...}: Timing.timing) msg = tracing ctxt ("found a " ^ msg ^ " in " ^ string_of_int (Time.toMilliseconds elapsed) ^ " ms") fun solve props ctxt = (case Solver_Data.get ctxt of NONE => error "bad Argo solver context" | SOME {names, types, terms, argo} => let val _ = with_tracing ctxt (trace_props props) val (es, (names', types', terms')) = fold_map (translate ctxt) props (names, types, terms) val _ = tracing ctxt "starting the prover" in (case Timing.timing (raw_solve es) argo of (time, Proof proof) => let val _ = trace_result ctxt time "proof" in (Proof (terms', proof), Solver_Data.put NONE ctxt) end | (time, Model argo') => let val _ = trace_result ctxt time "model" val model = value_of terms' (Argo_Solver.model_of argo') in (Model model, Solver_Data.put (SOME (mk_data names' types' terms' argo')) ctxt) end) end) (* reverse translation *) structure Contab = Table(type key = string * Argo_Expr.typ val ord = Argo_Expr.con_ord) fun mk_nary f ts = uncurry (fold_rev (curry f)) (split_last ts) fun mk_nary' d _ [] = d | mk_nary' _ f ts = mk_nary f ts fun mk_ite t1 t2 t3 = - let - val T = Term.fastype_of t2 - val ite = Const (\<^const_name>\HOL.If\, [\<^typ>\HOL.bool\, T, T] ---> T) - in ite $ t1 $ t2 $ t3 end + let val T = Term.fastype_of t2 + in \<^Const>\If T for t1 t2 t3\ end -fun term_of _ (Argo_Expr.E (Argo_Expr.True, _)) = \<^const>\HOL.True\ - | term_of _ (Argo_Expr.E (Argo_Expr.False, _)) = \<^const>\HOL.False\ +fun term_of _ (Argo_Expr.E (Argo_Expr.True, _)) = \<^Const>\True\ + | term_of _ (Argo_Expr.E (Argo_Expr.False, _)) = \<^Const>\False\ | term_of cx (Argo_Expr.E (Argo_Expr.Not, [e])) = HOLogic.mk_not (term_of cx e) | term_of cx (Argo_Expr.E (Argo_Expr.And, es)) = - mk_nary' \<^const>\HOL.True\ HOLogic.mk_conj (map (term_of cx) es) + mk_nary' \<^Const>\True\ HOLogic.mk_conj (map (term_of cx) es) | term_of cx (Argo_Expr.E (Argo_Expr.Or, es)) = - mk_nary' \<^const>\HOL.False\ HOLogic.mk_disj (map (term_of cx) es) + mk_nary' \<^Const>\False\ HOLogic.mk_disj (map (term_of cx) es) | term_of cx (Argo_Expr.E (Argo_Expr.Imp, [e1, e2])) = HOLogic.mk_imp (term_of cx e1, term_of cx e2) | term_of cx (Argo_Expr.E (Argo_Expr.Iff, [e1, e2])) = HOLogic.mk_eq (term_of cx e1, term_of cx e2) | term_of cx (Argo_Expr.E (Argo_Expr.Ite, [e1, e2, e3])) = mk_ite (term_of cx e1) (term_of cx e2) (term_of cx e3) | term_of cx (Argo_Expr.E (Argo_Expr.Eq, [e1, e2])) = HOLogic.mk_eq (term_of cx e1, term_of cx e2) | term_of cx (Argo_Expr.E (Argo_Expr.App, [e1, e2])) = term_of cx e1 $ term_of cx e2 | term_of (_, cons) (Argo_Expr.E (Argo_Expr.Con (c as (n, _)), _)) = (case Contab.lookup cons c of SOME t => t | NONE => error ("Unexpected expression named " ^ quote n)) | term_of (cx as (ctxt, _)) e = (case ext_term_of ctxt (term_of cx) e of SOME t => t | NONE => raise Fail "bad expression") -fun as_prop ct = Thm.apply \<^cterm>\HOL.Trueprop\ ct +fun as_prop ct = Thm.apply \<^cterm>\Trueprop\ ct fun cterm_of ctxt cons e = Thm.cterm_of ctxt (term_of (ctxt, cons) e) fun cprop_of ctxt cons e = as_prop (cterm_of ctxt cons e) (* generic proof tools *) fun discharge thm rule = thm INCR_COMP rule fun discharge2 thm1 thm2 rule = discharge thm2 (discharge thm1 rule) fun discharges thm rules = [thm] RL rules fun under_assumption f ct = let val cprop = as_prop ct in Thm.implies_intr cprop (f (Thm.assume cprop)) end fun instantiate cv ct = Thm.instantiate (TVars.empty, Vars.make [(Term.dest_Var (Thm.term_of cv), ct)]) (* proof replay for tautologies *) fun prove_taut ctxt ns t = Goal.prove ctxt ns [] (HOLogic.mk_Trueprop t) (fn {context, ...} => HEADGOAL (Classical.fast_tac context)) fun with_frees ctxt n mk = let val ns = map (fn i => "P" ^ string_of_int i) (0 upto (n - 1)) - val ts = map (Free o rpair \<^typ>\bool\) ns + val ts = map (Free o rpair \<^Type>\bool\) ns val t = mk_nary HOLogic.mk_disj (mk ts) in prove_taut ctxt ns t end fun taut_and1_term ts = mk_nary HOLogic.mk_conj ts :: map HOLogic.mk_not ts fun taut_and2_term i ts = [HOLogic.mk_not (mk_nary HOLogic.mk_conj ts), nth ts i] fun taut_or1_term i ts = [mk_nary HOLogic.mk_disj ts, HOLogic.mk_not (nth ts i)] fun taut_or2_term ts = HOLogic.mk_not (mk_nary HOLogic.mk_disj ts) :: ts val iff_1_taut = @{lemma "P = Q \ P \ Q" by fast} val iff_2_taut = @{lemma "P = Q \ (\P) \ (\Q)" by fast} val iff_3_taut = @{lemma "\(P = Q) \ (\P) \ Q" by fast} val iff_4_taut = @{lemma "\(P = Q) \ P \ (\Q)" by fast} val ite_then_taut = @{lemma "\P \ (if P then t else u) = t" by auto} val ite_else_taut = @{lemma "P \ (if P then t else u) = u" by auto} fun taut_rule_of ctxt (Argo_Proof.Taut_And_1 n) = with_frees ctxt n taut_and1_term | taut_rule_of ctxt (Argo_Proof.Taut_And_2 (i, n)) = with_frees ctxt n (taut_and2_term i) | taut_rule_of ctxt (Argo_Proof.Taut_Or_1 (i, n)) = with_frees ctxt n (taut_or1_term i) | taut_rule_of ctxt (Argo_Proof.Taut_Or_2 n) = with_frees ctxt n taut_or2_term | taut_rule_of _ Argo_Proof.Taut_Iff_1 = iff_1_taut | taut_rule_of _ Argo_Proof.Taut_Iff_2 = iff_2_taut | taut_rule_of _ Argo_Proof.Taut_Iff_3 = iff_3_taut | taut_rule_of _ Argo_Proof.Taut_Iff_4 = iff_4_taut | taut_rule_of _ Argo_Proof.Taut_Ite_Then = ite_then_taut | taut_rule_of _ Argo_Proof.Taut_Ite_Else = ite_else_taut fun replay_taut ctxt k ct = let val rule = taut_rule_of ctxt k in Thm.instantiate (Thm.match (Thm.cprop_of rule, ct)) rule end (* proof replay for conjunct extraction *) fun replay_conjunct 0 1 thm = thm | replay_conjunct 0 _ thm = discharge thm @{thm conjunct1} | replay_conjunct 1 2 thm = discharge thm @{thm conjunct2} | replay_conjunct i n thm = replay_conjunct (i - 1) (n - 1) (discharge thm @{thm conjunct2}) (* proof replay for rewrite steps *) fun mk_rewr thm = thm RS @{thm eq_reflection} fun not_nary_conv rule i ct = if i > 1 then (Conv.rewr_conv rule then_conv Conv.arg_conv (not_nary_conv rule (i - 1))) ct else Conv.all_conv ct val flatten_and_thm = @{lemma "(P1 \ P2) \ P3 \ P1 \ P2 \ P3" by simp} val flatten_or_thm = @{lemma "(P1 \ P2) \ P3 \ P1 \ P2 \ P3" by simp} fun flatten_conv cv rule ct = ( Conv.try_conv (Conv.arg_conv cv) then_conv Conv.try_conv (Conv.rewr_conv rule then_conv cv)) ct fun flat_conj_conv ct = (case Thm.term_of ct of - \<^const>\HOL.conj\ $ _ $ _ => flatten_conv flat_conj_conv flatten_and_thm ct + \<^Const_>\conj for _ _\ => flatten_conv flat_conj_conv flatten_and_thm ct | _ => Conv.all_conv ct) fun flat_disj_conv ct = (case Thm.term_of ct of - \<^const>\HOL.disj\ $ _ $ _ => flatten_conv flat_disj_conv flatten_or_thm ct + \<^Const_>\disj for _ _\ => flatten_conv flat_disj_conv flatten_or_thm ct | _ => Conv.all_conv ct) fun explode rule1 rule2 thm = explode rule1 rule2 (thm RS rule1) @ explode rule1 rule2 (thm RS rule2) handle THM _ => [thm] val explode_conj = explode @{thm conjunct1} @{thm conjunct2} val explode_ndis = explode @{lemma "\(P \ Q) \ \P" by auto} @{lemma "\(P \ Q) \ \Q" by auto} fun pick_false i thms = nth thms i fun pick_dual rule (i1, i2) thms = rule OF [nth thms i1, nth thms i2] handle THM _ => rule OF [nth thms i2, nth thms i1] val pick_dual_conj = pick_dual @{lemma "\P \ P \ False" by auto} val pick_dual_ndis = pick_dual @{lemma "\P \ P \ \True" by auto} fun join thm0 rule is thms = let val l = length thms val thms' = fold (fn i => cons (if 0 <= i andalso i < l then nth thms i else thm0)) is [] in fold (fn thm => fn thm' => discharge2 thm thm' rule) (tl thms') (hd thms') end val join_conj = join @{lemma "True" by auto} @{lemma "P \ Q \ P \ Q" by auto} val join_ndis = join @{lemma "\False" by auto} @{lemma "\P \ \Q \ \(P \ Q)" by auto} val false_thm = @{lemma "False \ P" by auto} val ntrue_thm = @{lemma "\True \ P" by auto} val iff_conj_thm = @{lemma "(P \ Q) \ (Q \ P) \ P = Q" by auto} val iff_ndis_thm = @{lemma "(\P \ \Q) \ (\Q \ \P) \ P = Q" by auto} fun iff_intro rule lf rf ct = let val lhs = under_assumption lf ct val rhs = rf (Thm.dest_arg (snd (Thm.dest_implies (Thm.cprop_of lhs)))) in mk_rewr (discharge2 lhs rhs rule) end fun with_conj f g ct = iff_intro iff_conj_thm (f o explode_conj) g ct -fun with_ndis f g ct = iff_intro iff_ndis_thm (f o explode_ndis) g (Thm.apply \<^cterm>\HOL.Not\ ct) +fun with_ndis f g ct = iff_intro iff_ndis_thm (f o explode_ndis) g (Thm.apply \<^cterm>\Not\ ct) fun swap_indices n iss = map (fn i => find_index (fn is => member (op =) is i) iss) (0 upto (n - 1)) fun sort_nary w f g (n, iss) = w (f (map hd iss)) (under_assumption (f (swap_indices n iss) o g)) val sort_conj = sort_nary with_conj join_conj explode_conj val sort_ndis = sort_nary with_ndis join_ndis explode_ndis val not_true_thm = mk_rewr @{lemma "(\True) = False" by auto} val not_false_thm = mk_rewr @{lemma "(\False) = True" by auto} val not_not_thm = mk_rewr @{lemma "(\\P) = P" by auto} val not_and_thm = mk_rewr @{lemma "(\(P \ Q)) = (\P \ \Q)" by auto} val not_or_thm = mk_rewr @{lemma "(\(P \ Q)) = (\P \ \Q)" by auto} val not_iff_thms = map mk_rewr @{lemma "(\((\P) = Q)) = (P = Q)" "(\(P = (\Q))) = (P = Q)" "(\(P = Q)) = ((\P) = Q)" by auto} val iff_true_thms = map mk_rewr @{lemma "(True = P) = P" "(P = True) = P" by auto} val iff_false_thms = map mk_rewr @{lemma "(False = P) = (\P)" "(P = False) = (\P)" by auto} val iff_not_not_thm = mk_rewr @{lemma "((\P) = (\Q)) = (P = Q)" by auto} val iff_refl_thm = mk_rewr @{lemma "(P = P) = True" by auto} val iff_symm_thm = mk_rewr @{lemma "(P = Q) = (Q = P)" by auto} val iff_dual_thms = map mk_rewr @{lemma "(P = (\P)) = False" "((\P) = P) = False" by auto} val imp_thm = mk_rewr @{lemma "(P \ Q) = (\P \ Q)" by auto} val ite_prop_thm = mk_rewr @{lemma "(If P Q R) = ((\P \ Q) \ (P \ R) \ (Q \ R))" by auto} val ite_true_thm = mk_rewr @{lemma "(If True t u) = t" by auto} val ite_false_thm = mk_rewr @{lemma "(If False t u) = u" by auto} val ite_eq_thm = mk_rewr @{lemma "(If P t t) = t" by auto} val eq_refl_thm = mk_rewr @{lemma "(t = t) = True" by auto} val eq_symm_thm = mk_rewr @{lemma "(t1 = t2) = (t2 = t1)" by auto} fun replay_rewr _ Argo_Proof.Rewr_Not_True = Conv.rewr_conv not_true_thm | replay_rewr _ Argo_Proof.Rewr_Not_False = Conv.rewr_conv not_false_thm | replay_rewr _ Argo_Proof.Rewr_Not_Not = Conv.rewr_conv not_not_thm | replay_rewr _ (Argo_Proof.Rewr_Not_And i) = not_nary_conv not_and_thm i | replay_rewr _ (Argo_Proof.Rewr_Not_Or i) = not_nary_conv not_or_thm i | replay_rewr _ Argo_Proof.Rewr_Not_Iff = Conv.rewrs_conv not_iff_thms | replay_rewr _ (Argo_Proof.Rewr_And_False i) = with_conj (pick_false i) (K false_thm) | replay_rewr _ (Argo_Proof.Rewr_And_Dual ip) = with_conj (pick_dual_conj ip) (K false_thm) | replay_rewr _ (Argo_Proof.Rewr_And_Sort is) = flat_conj_conv then_conv sort_conj is | replay_rewr _ (Argo_Proof.Rewr_Or_True i) = with_ndis (pick_false i) (K ntrue_thm) | replay_rewr _ (Argo_Proof.Rewr_Or_Dual ip) = with_ndis (pick_dual_ndis ip) (K ntrue_thm) | replay_rewr _ (Argo_Proof.Rewr_Or_Sort is) = flat_disj_conv then_conv sort_ndis is | replay_rewr _ Argo_Proof.Rewr_Iff_True = Conv.rewrs_conv iff_true_thms | replay_rewr _ Argo_Proof.Rewr_Iff_False = Conv.rewrs_conv iff_false_thms | replay_rewr _ Argo_Proof.Rewr_Iff_Not_Not = Conv.rewr_conv iff_not_not_thm | replay_rewr _ Argo_Proof.Rewr_Iff_Refl = Conv.rewr_conv iff_refl_thm | replay_rewr _ Argo_Proof.Rewr_Iff_Symm = Conv.rewr_conv iff_symm_thm | replay_rewr _ Argo_Proof.Rewr_Iff_Dual = Conv.rewrs_conv iff_dual_thms | replay_rewr _ Argo_Proof.Rewr_Imp = Conv.rewr_conv imp_thm | replay_rewr _ Argo_Proof.Rewr_Ite_Prop = Conv.rewr_conv ite_prop_thm | replay_rewr _ Argo_Proof.Rewr_Ite_True = Conv.rewr_conv ite_true_thm | replay_rewr _ Argo_Proof.Rewr_Ite_False = Conv.rewr_conv ite_false_thm | replay_rewr _ Argo_Proof.Rewr_Ite_Eq = Conv.rewr_conv ite_eq_thm | replay_rewr _ Argo_Proof.Rewr_Eq_Refl = Conv.rewr_conv eq_refl_thm | replay_rewr _ Argo_Proof.Rewr_Eq_Symm = Conv.rewr_conv eq_symm_thm | replay_rewr ctxt r = ext_replay_rewr ctxt r fun binop_conv cv1 cv2 = Conv.combination_conv (Conv.arg_conv cv1) cv2 fun replay_conv _ Argo_Proof.Keep_Conv ct = Conv.all_conv ct | replay_conv ctxt (Argo_Proof.Then_Conv (c1, c2)) ct = (replay_conv ctxt c1 then_conv replay_conv ctxt c2) ct | replay_conv ctxt (Argo_Proof.Args_Conv (Argo_Expr.App, [c1, c2])) ct = Conv.combination_conv (replay_conv ctxt c1) (replay_conv ctxt c2) ct | replay_conv ctxt (Argo_Proof.Args_Conv (_, cs)) ct = replay_args_conv ctxt cs ct | replay_conv ctxt (Argo_Proof.Rewr_Conv r) ct = replay_rewr ctxt r ct and replay_args_conv _ [] ct = Conv.all_conv ct | replay_args_conv ctxt [c] ct = Conv.arg_conv (replay_conv ctxt c) ct | replay_args_conv ctxt [c1, c2] ct = binop_conv (replay_conv ctxt c1) (replay_conv ctxt c2) ct | replay_args_conv ctxt (c :: cs) ct = (case Term.head_of (Thm.term_of ct) of - Const (\<^const_name>\HOL.If\, _) => + \<^Const_>\If _\ => let val (cs', c') = split_last cs in Conv.combination_conv (replay_args_conv ctxt (c :: cs')) (replay_conv ctxt c') ct end | _ => binop_conv (replay_conv ctxt c) (replay_args_conv ctxt cs) ct) fun replay_rewrite ctxt c thm = Conv.fconv_rule (HOLogic.Trueprop_conv (replay_conv ctxt c)) thm (* proof replay for clauses *) val prep_clause_rule = @{lemma "P \ \P \ False" by fast} val extract_lit_rule = @{lemma "(\(P \ Q) \ False) \ \P \ \Q \ False" by fast} fun add_lit i thm lits = let val ct = Thm.cprem_of thm 1 in (Thm.implies_elim thm (Thm.assume ct), (i, ct) :: lits) end fun extract_lits [] _ = error "Bad clause" | extract_lits [i] (thm, lits) = add_lit i thm lits | extract_lits (i :: is) (thm, lits) = extract_lits is (add_lit i (discharge thm extract_lit_rule) lits) fun lit_ord ((l1, _), (l2, _)) = int_ord (abs l1, abs l2) fun replay_with_lits [] thm lits = (thm, lits) | replay_with_lits is thm lits = extract_lits is (discharge thm prep_clause_rule, lits) ||> Ord_List.make lit_ord fun replay_clause is thm = replay_with_lits is thm [] (* proof replay for unit resolution *) val unit_rule = @{lemma "(P \ False) \ (\P \ False) \ False" by fast} val unit_rule_var = Thm.dest_arg (Thm.dest_arg1 (Thm.cprem_of unit_rule 1)) -val bogus_ct = \<^cterm>\HOL.True\ +val bogus_ct = \<^cterm>\True\ fun replay_unit_res lit (pthm, plits) (nthm, nlits) = let val plit = the (AList.lookup (op =) plits lit) val nlit = the (AList.lookup (op =) nlits (~lit)) val prune = Ord_List.remove lit_ord (lit, bogus_ct) in unit_rule |> instantiate unit_rule_var (Thm.dest_arg plit) |> Thm.elim_implies (Thm.implies_intr plit pthm) |> Thm.elim_implies (Thm.implies_intr nlit nthm) |> rpair (Ord_List.union lit_ord (prune nlits) (prune plits)) end (* proof replay for hypothesis *) val dneg_rule = @{lemma "\\P \ P" by auto} fun replay_hyp i ct = if i < 0 then (Thm.assume ct, [(~i, ct)]) else - let val cu = as_prop (Thm.apply \<^cterm>\HOL.Not\ (Thm.apply \<^cterm>\HOL.Not\ (Thm.dest_arg ct))) + let val cu = as_prop (Thm.apply \<^cterm>\Not\ (Thm.apply \<^cterm>\Not\ (Thm.dest_arg ct))) in (discharge (Thm.assume cu) dneg_rule, [(~i, cu)]) end (* proof replay for lemma *) fun replay_lemma is (thm, lits) = replay_with_lits is thm lits (* proof replay for reflexivity *) val refl_rule = @{thm refl} val refl_rule_var = Thm.dest_arg1 (Thm.dest_arg (Thm.cprop_of refl_rule)) fun replay_refl ct = Thm.instantiate (Thm.match (refl_rule_var, ct)) refl_rule (* proof replay for symmetry *) val symm_rules = @{lemma "a = b ==> b = a" "\(a = b) \ \(b = a)" by simp_all} fun replay_symm thm = hd (discharges thm symm_rules) (* proof replay for transitivity *) val trans_rules = @{lemma "\(a = b) \ b = c \ \(a = c)" "a = b \ \(b = c) \ \(a = c)" "a = b \ b = c ==> a = c" by simp_all} fun replay_trans thm1 thm2 = hd (discharges thm2 (discharges thm1 trans_rules)) (* proof replay for congruence *) fun replay_cong thm1 thm2 = discharge thm2 (discharge thm1 @{thm cong}) (* proof replay for substitution *) val subst_rule1 = @{lemma "\(p a) \ p = q \ a = b \ \(q b)" by simp} val subst_rule2 = @{lemma "p a \ p = q \ a = b \ q b" by simp} fun replay_subst thm1 thm2 thm3 = subst_rule1 OF [thm1, thm2, thm3] handle THM _ => subst_rule2 OF [thm1, thm2, thm3] (* proof replay *) structure Thm_Cache = Table(type key = Argo_Proof.proof_id val ord = Argo_Proof.proof_id_ord) val unclausify_rule1 = @{lemma "(\P \ False) \ P" by auto} val unclausify_rule2 = @{lemma "(P \ False) \ \P" by auto} fun unclausify (thm, lits) ls = (case (Thm.prop_of thm, lits) of - (\<^const>\HOL.Trueprop\ $ \<^const>\HOL.False\, [(_, ct)]) => + (\<^Const_>\Trueprop for \\<^Const_>\False\\\, [(_, ct)]) => let val thm = Thm.implies_intr ct thm in (discharge thm unclausify_rule1 handle THM _ => discharge thm unclausify_rule2, ls) end | _ => (thm, Ord_List.union lit_ord lits ls)) fun with_thms f tps = fold_map unclausify tps [] |>> f fun bad_premises () = raise Fail "bad number of premises" fun with_thms1 f = with_thms (fn [thm] => f thm | _ => bad_premises ()) fun with_thms2 f = with_thms (fn [thm1, thm2] => f thm1 thm2 | _ => bad_premises ()) fun with_thms3 f = with_thms (fn [thm1, thm2, thm3] => f thm1 thm2 thm3 | _ => bad_premises ()) fun replay_rule (ctxt, cons, facts) prems rule = (case rule of Argo_Proof.Axiom i => (nth facts i, []) | Argo_Proof.Taut (k, concl) => (replay_taut ctxt k (cprop_of ctxt cons concl), []) | Argo_Proof.Conjunct (i, n) => with_thms1 (replay_conjunct i n) prems | Argo_Proof.Rewrite c => with_thms1 (replay_rewrite ctxt c) prems | Argo_Proof.Clause is => replay_clause is (fst (hd prems)) | Argo_Proof.Unit_Res i => replay_unit_res i (hd prems) (hd (tl prems)) | Argo_Proof.Hyp (i, concl) => replay_hyp i (cprop_of ctxt cons concl) | Argo_Proof.Lemma is => replay_lemma is (hd prems) | Argo_Proof.Refl concl => (replay_refl (cterm_of ctxt cons concl), []) | Argo_Proof.Symm => with_thms1 replay_symm prems | Argo_Proof.Trans => with_thms2 replay_trans prems | Argo_Proof.Cong => with_thms2 replay_cong prems | Argo_Proof.Subst => with_thms3 replay_subst prems | _ => with_thms (ext_replay (cprop_of ctxt cons) ctxt rule) prems) fun with_cache f proof thm_cache = (case Thm_Cache.lookup thm_cache (Argo_Proof.id_of proof) of SOME thm => (thm, thm_cache) | NONE => let val (thm, thm_cache') = f proof thm_cache in (thm, Thm_Cache.update (Argo_Proof.id_of proof, thm) thm_cache') end) fun trace_step ctxt proof_id rule proofs = with_full_tracing ctxt (fn ctxt' => let val id = Argo_Proof.string_of_proof_id proof_id val ids = map (Argo_Proof.string_of_proof_id o Argo_Proof.id_of) proofs val rule_string = Argo_Proof.string_of_rule rule in full_tracing ctxt' (" " ^ id ^ " <- " ^ space_implode " " ids ^ " . " ^ rule_string) end) fun replay_bottom_up (env as (ctxt, _, _)) proof thm_cache = let val (proof_id, rule, proofs) = Argo_Proof.dest proof val (prems, thm_cache) = fold_map (with_cache (replay_bottom_up env)) proofs thm_cache val _ = trace_step ctxt proof_id rule proofs in (replay_rule env prems rule, thm_cache) end fun replay_proof env proof = with_cache (replay_bottom_up env) proof Thm_Cache.empty fun replay ctxt terms facts proof = let val env = (ctxt, Termtab.fold (Contab.update o swap) terms Contab.empty, facts) val _ = tracing ctxt "replaying the proof" val ({elapsed=t, ...}, ((thm, _), _)) = Timing.timing (replay_proof env) proof val _ = tracing ctxt ("replayed the proof in " ^ string_of_int (Time.toMilliseconds t) ^ " ms") in thm end (* normalizing goals *) fun instantiate_elim_rule thm = let val ct = Drule.strip_imp_concl (Thm.cprop_of thm) in (case Thm.term_of ct of - \<^const>\HOL.Trueprop\ $ Var (_, \<^typ>\HOL.bool\) => - instantiate (Thm.dest_arg ct) \<^cterm>\HOL.False\ thm - | Var _ => instantiate ct \<^cprop>\HOL.False\ thm + \<^Const_>\Trueprop for \Var (_, \<^Type>\bool\)\\ => + instantiate (Thm.dest_arg ct) \<^cterm>\False\ thm + | Var _ => instantiate ct \<^cprop>\False\ thm | _ => thm) end fun atomize_conv ctxt ct = (case Thm.term_of ct of - \<^const>\HOL.Trueprop\ $ _ => Conv.all_conv - | \<^const>\Pure.imp\ $ _ $ _ => + \<^Const_>\Trueprop for _\ => Conv.all_conv + | \<^Const_>\Pure.imp for _ _\ => Conv.binop_conv (atomize_conv ctxt) then_conv Conv.rewr_conv @{thm atomize_imp} - | Const (\<^const_name>\Pure.eq\, _) $ _ $ _ => + | \<^Const>\Pure.eq _ for _ _\ => Conv.binop_conv (atomize_conv ctxt) then_conv Conv.rewr_conv @{thm atomize_eq} - | Const (\<^const_name>\Pure.all\, _) $ Abs _ => + | \<^Const>\Pure.all _ for \Abs _\\ => Conv.binder_conv (atomize_conv o snd) ctxt then_conv Conv.rewr_conv @{thm atomize_all} | _ => Conv.all_conv) ct fun normalize ctxt thm = thm |> instantiate_elim_rule |> Conv.fconv_rule (Thm.beta_conversion true then_conv Thm.eta_conversion) |> Thm.forall_intr_vars |> Conv.fconv_rule (atomize_conv ctxt) (* prover, tactic and method *) datatype result = Satisfiable of term -> bool option | Unsatisfiable fun check props ctxt = (case solve props ctxt of (Proof _, ctxt') => (Unsatisfiable, ctxt') | (Model model, ctxt') => (Satisfiable model, ctxt')) fun prove thms ctxt = let val thms' = map (normalize ctxt) thms in (case solve (map Thm.prop_of thms') ctxt of (Model _, ctxt') => (NONE, ctxt') | (Proof (terms, proof), ctxt') => (SOME (replay ctxt' terms thms' proof), ctxt')) end fun argo_tac ctxt thms = CONVERSION (Conv.params_conv ~1 (K (Conv.concl_conv ~1 (Conv.try_conv (Conv.rewr_conv @{thm atomize_eq})))) ctxt) THEN' Tactic.resolve_tac ctxt [@{thm ccontr}] THEN' Subgoal.FOCUS (fn {context, prems, ...} => (case with_timeout context (prove (thms @ prems)) context of (SOME thm, _) => Tactic.resolve_tac context [thm] 1 | (NONE, _) => Tactical.no_tac)) ctxt val _ = Theory.setup (Method.setup \<^binding>\argo\ (Scan.optional Attrib.thms [] >> (fn thms => fn ctxt => METHOD (fn facts => HEADGOAL (argo_tac ctxt (thms @ facts))))) "Applies the Argo prover") end