# HG changeset patch # User hoelzl # Date 1428831199 -7200 # Node ID 218fcc645d22e9a27b200cb574fd36f718d43547 # Parent 4b77fc0b3514416db52d15f56df3ad7f7e149911 move filters to their own theory diff -r 4b77fc0b3514 -r 218fcc645d22 src/HOL/Filter.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Filter.thy Sun Apr 12 11:33:19 2015 +0200 @@ -0,0 +1,966 @@ +(* 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" + +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 eventually_mono: + "(\x. P x \ Q x) \ eventually P F \ eventually Q F" + unfolding eventually_def + by (rule 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_Ball_finite: + assumes "finite A" and "\y\A. eventually (\x. P x y) net" + shows "eventually (\x. \y\A. P x y) net" +using assms by (induct set: finite, simp, simp add: eventually_conj) + +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_mp: + assumes "eventually (\x. P x \ Q x) F" + assumes "eventually (\x. P x) F" + shows "eventually (\x. Q x) F" +proof (rule eventually_mono) + show "\x. (P x \ Q x) \ P x \ Q x" by simp + show "eventually (\x. (P x \ Q x) \ P x) F" + using assms by (rule eventually_conj) +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_elim1: + assumes "eventually (\i. P i) F" + assumes "\i. P i \ Q i" + shows "eventually (\i. Q i) F" + using assms by (auto 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 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_elim1) + then show ?thesis by (auto elim: eventually_elim2) +qed + +ML {* + fun eventually_elim_tac ctxt facts = SUBGOAL_CASES (fn (goal, i) => + let + val mp_thms = facts RL @{thms eventually_rev_mp} + val raw_elim_thm = + (@{thm allI} RS @{thm always_eventually}) + |> fold (fn thm1 => fn thm2 => thm2 RS thm1) mp_thms + |> fold (fn _ => fn thm => @{thm impI} RS thm) facts + val cases_prop = Thm.prop_of (raw_elim_thm RS Goal.init (Thm.cterm_of ctxt goal)) + val cases = Rule_Cases.make_common ctxt cases_prop [(("elim", []), [])] + in + CASES cases (rtac raw_elim_thm i) + end) +*} + +method_setup eventually_elim = {* + Scan.succeed (fn ctxt => METHOD_CASES (HEADGOAL o eventually_elim_tac ctxt)) +*} "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, rule is_filter.intro) + apply (fast intro: eventually_True) + apply clarify + apply (intro exI conjI) + apply (erule (1) eventually_conj) + apply (erule (1) eventually_conj) + apply simp + apply auto + done + +lemma eventually_Sup: + "eventually P (Sup S) \ (\F\S. eventually P F)" + unfolding Sup_filter_def + apply (rule eventually_Abs_filter, rule 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 elim!: eventually_mono intro: 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 + +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) + +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 eventually_const: "\ trivial_limit net \ eventually (\x. P) net \ P" + by (cases P) (simp_all 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_elim1) + qed } + note eventually_F = this + + 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 (INF b:B. F b) \ (\X\B. finite X \ eventually P (INF b:X. F b))" + unfolding INF_def[of B] eventually_Inf[of P "F`B"] + by (metis Inf_image_eq 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 \ (INF b:X. F b) \ bot) \ (INF 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 (INF b:B. F b) \ (\b\B. eventually P (F b))" + unfolding INF_def by (subst eventually_Inf_base) 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) + apply (rule 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_sup: "filtermap f (sup F1 F2) = sup (filtermap f F1) (filtermap f F2)" + by (auto simp: filter_eq_iff eventually_filtermap eventually_sup) + +lemma filtermap_inf: "filtermap f (inf F1 F2) \ inf (filtermap f F1) (filtermap f F2)" + by (auto simp: le_filter_def eventually_filtermap eventually_inf) + +lemma filtermap_INF: "filtermap f (INF b:B. F b) \ (INF b:B. filtermap f (F b))" +proof - + { fix X :: "'c set" assume "finite X" + then have "filtermap f (INFIMUM X F) \ (INF b:X. filtermap f (F b))" + proof induct + case (insert x X) + have "filtermap f (INF a:insert x X. F a) \ inf (filtermap f (F x)) (filtermap f (INF a:X. F a))" + by (rule order_trans[OF _ filtermap_inf]) simp + also have "\ \ inf (filtermap f (F x)) (INF a:X. filtermap f (F a))" + by (intro inf_mono insert order_refl) + finally show ?case + by simp + qed simp } + then show ?thesis + unfolding le_filter_def eventually_filtermap + by (subst (1 2) eventually_INF) auto +qed +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_elim1) + +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 + apply safe + apply (erule_tac x="\x. x \ A" in allE) + apply (auto elim: eventually_elim1) + done + +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]: "(SUP i : I. principal (A i)) = principal (\i\I. A i)" + unfolding filter_eq_iff eventually_Sup SUP_def by (auto simp: eventually_principal) + +lemma INF_principal_finite: "finite X \ (INF 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 + +subsubsection {* Order filters *} + +definition at_top :: "('a::order) filter" + where "at_top = (INF k. principal {k ..})" + +lemma at_top_sub: "at_top = (INF 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_ge_at_top: + "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 (INF 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 "(INF 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_gt_at_top: + "eventually (\x. (c::_::unbounded_dense_linorder) < x) at_top" + unfolding eventually_at_top_dense by auto + +definition at_bot :: "('a::order) filter" + where "at_bot = (INF k. principal {.. k})" + +lemma at_bot_sub: "at_bot = (INF 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_le_at_bot: + "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}. \n (\N::'a. \nx. x < (c::_::unbounded_dense_linorder)) at_bot" + unfolding eventually_at_bot_dense by auto + +lemma trivial_limit_at_bot_linorder: "\ 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: "\ 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: + assumes "\x. c \ x \ P x" + shows "eventually P sequentially" +using assms by (auto simp: eventually_sequentially) + +lemma eventually_sequentially_seg: + "eventually (\n. P (n + k)) sequentially \ eventually P sequentially" + unfolding eventually_sequentially + apply safe + apply (rule_tac x="N + k" in exI) + apply rule + apply (erule_tac x="n - k" in allE) + apply auto [] + apply (rule_tac x=N in exI) + apply auto [] + done + + +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_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'" + apply (rule filterlim_cong[OF refl refl eq, THEN iffD1]) + apply (rule filterlim_mono[OF _ ord]) + apply fact + done + +lemma filtermap_mono_strong: "inj f \ filtermap f F \ filtermap f G \ F \ G" + apply (auto 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 filtermap_eq_strong: "inj f \ filtermap f F = filtermap f G \ F = G" + by (simp add: filtermap_mono_strong eq_iff) + +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_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 :> (INF 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 (INF i:I. F i). f x :> (INF j:J. G j)" + unfolding filterlim_def by (rule order_trans[OF filtermap_INF INF_mono]) + +lemma filterlim_base: + "(\m x. m \ J \ i m \ I) \ (\m x. m \ J \ x \ F (i m) \ f x \ G m) \ + LIM x (INF i:I. principal (F i)). f x :> (INF 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 (INF i:I. principal (F i)). f x :> INF 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 eventually_sequentially_Suc: "eventually (\i. P (Suc i)) sequentially \ eventually P sequentially" + unfolding eventually_sequentially by (metis Suc_le_D Suc_le_mono le_Suc_eq) + +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) (metis le_Suc_eq) + + +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_elim1) + +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_elim1[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_elim1) + +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" + apply (rule eventually_mono[rotated]) + using 1 by 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[rotated], 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_elim1) +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) + + +subsection {* Setup @{typ "'a filter"} for lifting and transfer *} + +context begin interpretation lifting_syntax . + +definition rel_filter :: "('a \ 'b \ bool) \ 'a filter \ 'b filter \ bool" +where "rel_filter R F G = ((R ===> op =) ===> op =) (Rep_filter F) (Rep_filter G)" + +lemma rel_filter_eventually: + "rel_filter R F G \ + ((R ===> op =) ===> op =) (\P. eventually P F) (\P. eventually P G)" +by(simp add: rel_filter_def eventually_def) + +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 . + +lemma Quotient_filter [quot_map]: + assumes Q: "Quotient R Abs Rep T" + shows "Quotient (rel_filter R) (filtermap Abs) (filtermap Rep) (rel_filter T)" +unfolding Quotient_alt_def +proof(intro conjI strip) + from Q have *: "\x y. T x y \ Abs x = y" + unfolding Quotient_alt_def by blast + + fix F G + assume "rel_filter T F G" + thus "filtermap Abs F = G" unfolding filter_eq_iff + by(auto simp add: eventually_filtermap rel_filter_eventually * rel_funI del: iffI elim!: rel_funD) +next + from Q have *: "\x. T (Rep x) x" unfolding Quotient_alt_def by blast + + fix F + show "rel_filter T (filtermap Rep F) F" + by(auto elim: rel_funD intro: * intro!: ext arg_cong[where f="\P. eventually P F"] rel_funI + del: iffI simp add: eventually_filtermap rel_filter_eventually) +qed(auto simp add: map_fun_def o_def eventually_filtermap filter_eq_iff fun_eq_iff rel_filter_eventually + fun_quotient[OF fun_quotient[OF Q identity_quotient] identity_quotient, unfolded Quotient_alt_def]) + +lemma eventually_parametric [transfer_rule]: + "((A ===> op =) ===> rel_filter A ===> op =) eventually eventually" +by(simp add: rel_fun_def rel_filter_eventually) + +lemma rel_filter_eq [relator_eq]: "rel_filter op = = op =" +by(auto simp add: rel_filter_eventually rel_fun_eq fun_eq_iff filter_eq_iff) + +lemma rel_filter_mono [relator_mono]: + "A \ B \ rel_filter A \ rel_filter B" +unfolding rel_filter_eventually[abs_def] +by(rule le_funI)+(intro fun_mono fun_mono[THEN le_funD, THEN le_funD] order.refl) + +lemma rel_filter_conversep [simp]: "rel_filter A\\ = (rel_filter A)\\" +by(auto simp add: rel_filter_eventually fun_eq_iff rel_fun_def) + +lemma is_filter_parametric_aux: + assumes "is_filter F" + assumes [transfer_rule]: "bi_total A" "bi_unique A" + and [transfer_rule]: "((A ===> op =) ===> op =) F G" + shows "is_filter G" +proof - + interpret is_filter F by fact + show ?thesis + proof + have "F (\_. True) = G (\x. True)" by transfer_prover + thus "G (\x. True)" by(simp add: True) + next + fix P' Q' + assume "G P'" "G Q'" + moreover + from bi_total_fun[OF `bi_unique A` bi_total_eq, unfolded bi_total_def] + obtain P Q where [transfer_rule]: "(A ===> op =) P P'" "(A ===> op =) Q Q'" by blast + have "F P = G P'" "F Q = G Q'" by transfer_prover+ + ultimately have "F (\x. P x \ Q x)" by(simp add: conj) + moreover have "F (\x. P x \ Q x) = G (\x. P' x \ Q' x)" by transfer_prover + ultimately show "G (\x. P' x \ Q' x)" by simp + next + fix P' Q' + assume "\x. P' x \ Q' x" "G P'" + moreover + from bi_total_fun[OF `bi_unique A` bi_total_eq, unfolded bi_total_def] + obtain P Q where [transfer_rule]: "(A ===> op =) P P'" "(A ===> op =) Q Q'" by blast + have "F P = G P'" by transfer_prover + moreover have "(\x. P x \ Q x) \ (\x. P' x \ Q' x)" by transfer_prover + ultimately have "F Q" by(simp add: mono) + moreover have "F Q = G Q'" by transfer_prover + ultimately show "G Q'" by simp + qed +qed + +lemma is_filter_parametric [transfer_rule]: + "\ bi_total A; bi_unique A \ + \ (((A ===> op =) ===> op =) ===> op =) is_filter is_filter" +apply(rule rel_funI) +apply(rule iffI) + apply(erule (3) is_filter_parametric_aux) +apply(erule is_filter_parametric_aux[where A="conversep A"]) +apply(auto simp add: rel_fun_def) +done + +lemma left_total_rel_filter [transfer_rule]: + assumes [transfer_rule]: "bi_total A" "bi_unique A" + shows "left_total (rel_filter A)" +proof(rule left_totalI) + fix F :: "'a filter" + from bi_total_fun[OF bi_unique_fun[OF `bi_total A` bi_unique_eq] bi_total_eq] + obtain G where [transfer_rule]: "((A ===> op =) ===> op =) (\P. eventually P F) G" + unfolding bi_total_def by blast + moreover have "is_filter (\P. eventually P F) \ is_filter G" by transfer_prover + hence "is_filter G" by(simp add: eventually_def is_filter_Rep_filter) + ultimately have "rel_filter A F (Abs_filter G)" + by(simp add: rel_filter_eventually eventually_Abs_filter) + thus "\G. rel_filter A F G" .. +qed + +lemma right_total_rel_filter [transfer_rule]: + "\ bi_total A; bi_unique A \ \ right_total (rel_filter A)" +using left_total_rel_filter[of "A\\"] by simp + +lemma bi_total_rel_filter [transfer_rule]: + assumes "bi_total A" "bi_unique A" + shows "bi_total (rel_filter A)" +unfolding bi_total_alt_def using assms +by(simp add: left_total_rel_filter right_total_rel_filter) + +lemma left_unique_rel_filter [transfer_rule]: + assumes "left_unique A" + shows "left_unique (rel_filter A)" +proof(rule left_uniqueI) + fix F F' G + assume [transfer_rule]: "rel_filter A F G" "rel_filter A F' G" + show "F = F'" + unfolding filter_eq_iff + proof + fix P :: "'a \ bool" + obtain P' where [transfer_rule]: "(A ===> op =) P P'" + using left_total_fun[OF assms left_total_eq] unfolding left_total_def by blast + have "eventually P F = eventually P' G" + and "eventually P F' = eventually P' G" by transfer_prover+ + thus "eventually P F = eventually P F'" by simp + qed +qed + +lemma right_unique_rel_filter [transfer_rule]: + "right_unique A \ right_unique (rel_filter A)" +using left_unique_rel_filter[of "A\\"] by simp + +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 top_filter_parametric [transfer_rule]: + "bi_total A \ (rel_filter A) top top" +by(simp add: rel_filter_eventually All_transfer) + +lemma bot_filter_parametric [transfer_rule]: "(rel_filter A) bot bot" +by(simp add: rel_filter_eventually rel_fun_def) + +lemma sup_filter_parametric [transfer_rule]: + "(rel_filter A ===> rel_filter A ===> rel_filter A) sup sup" +by(fastforce simp add: rel_filter_eventually[abs_def] eventually_sup dest: rel_funD) + +lemma Sup_filter_parametric [transfer_rule]: + "(rel_set (rel_filter A) ===> rel_filter A) Sup Sup" +proof(rule rel_funI) + fix S T + assume [transfer_rule]: "rel_set (rel_filter A) S T" + show "rel_filter A (Sup S) (Sup T)" + by(simp add: rel_filter_eventually eventually_Sup) transfer_prover +qed + +lemma principal_parametric [transfer_rule]: + "(rel_set A ===> rel_filter A) principal principal" +proof(rule rel_funI) + fix S S' + assume [transfer_rule]: "rel_set A S S'" + show "rel_filter A (principal S) (principal S')" + by(simp add: rel_filter_eventually eventually_principal) transfer_prover +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 ===> op =) op \ op \" +unfolding le_filter_def[abs_def] by transfer_prover + +lemma less_filter_parametric [transfer_rule]: + "(rel_filter A ===> rel_filter A ===> op =) op < op <" +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 + +end \ No newline at end of file diff -r 4b77fc0b3514 -r 218fcc645d22 src/HOL/Main.thy --- a/src/HOL/Main.thy Sun Apr 12 16:04:53 2015 +0200 +++ b/src/HOL/Main.thy Sun Apr 12 11:33:19 2015 +0200 @@ -1,7 +1,7 @@ section {* Main HOL *} theory Main -imports Predicate_Compile Quickcheck_Narrowing Extraction Coinduction Nitpick BNF_Greatest_Fixpoint +imports Predicate_Compile Quickcheck_Narrowing Extraction Coinduction Nitpick BNF_Greatest_Fixpoint Filter begin text {* diff -r 4b77fc0b3514 -r 218fcc645d22 src/HOL/NSA/Filter.thy --- a/src/HOL/NSA/Filter.thy Sun Apr 12 16:04:53 2015 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,413 +0,0 @@ -(* Title: HOL/NSA/Filter.thy - Author: Jacques D. Fleuriot, University of Cambridge - Author: Lawrence C Paulson - Author: Brian Huffman -*) - -section {* Filters and Ultrafilters *} - -theory Filter -imports "~~/src/HOL/Library/Infinite_Set" -begin - -subsection {* Definitions and basic properties *} - -subsubsection {* Filters *} - -locale filter = - fixes F :: "'a set set" - assumes UNIV [iff]: "UNIV \ F" - assumes empty [iff]: "{} \ F" - assumes Int: "\u \ F; v \ F\ \ u \ v \ F" - assumes subset: "\u \ F; u \ v\ \ v \ F" -begin - -lemma memD: "A \ F \ - A \ F" -proof - assume "A \ F" and "- A \ F" - hence "A \ (- A) \ F" by (rule Int) - thus "False" by simp -qed - -lemma not_memI: "- A \ F \ A \ F" -by (drule memD, simp) - -lemma Int_iff: "(x \ y \ F) = (x \ F \ y \ F)" -by (auto elim: subset intro: Int) - -end - -subsubsection {* Ultrafilters *} - -locale ultrafilter = filter + - assumes ultra: "A \ F \ - A \ F" -begin - -lemma memI: "- A \ F \ A \ F" -using ultra [of A] by simp - -lemma not_memD: "A \ F \ - A \ F" -by (rule memI, simp) - -lemma not_mem_iff: "(A \ F) = (- A \ F)" -by (rule iffI [OF not_memD not_memI]) - -lemma Compl_iff: "(- A \ F) = (A \ F)" -by (rule iffI [OF not_memI not_memD]) - -lemma Un_iff: "(x \ y \ F) = (x \ F \ y \ F)" - apply (rule iffI) - apply (erule contrapos_pp) - apply (simp add: Int_iff not_mem_iff) - apply (auto elim: subset) -done - -end - -subsubsection {* Free Ultrafilters *} - -locale freeultrafilter = ultrafilter + - assumes infinite: "A \ F \ infinite A" -begin - -lemma finite: "finite A \ A \ F" -by (erule contrapos_pn, erule infinite) - -lemma singleton: "{x} \ F" -by (rule finite, simp) - -lemma insert_iff [simp]: "(insert x A \ F) = (A \ F)" -apply (subst insert_is_Un) -apply (subst Un_iff) -apply (simp add: singleton) -done - -lemma filter: "filter F" .. - -lemma ultrafilter: "ultrafilter F" .. - -end - -subsection {* Collect properties *} - -lemma (in filter) Collect_ex: - "({n. \x. P n x} \ F) = (\X. {n. P n (X n)} \ F)" -proof - assume "{n. \x. P n x} \ F" - hence "{n. P n (SOME x. P n x)} \ F" - by (auto elim: someI subset) - thus "\X. {n. P n (X n)} \ F" by fast -next - show "\X. {n. P n (X n)} \ F \ {n. \x. P n x} \ F" - by (auto elim: subset) -qed - -lemma (in filter) Collect_conj: - "({n. P n \ Q n} \ F) = ({n. P n} \ F \ {n. Q n} \ F)" -by (subst Collect_conj_eq, rule Int_iff) - -lemma (in ultrafilter) Collect_not: - "({n. \ P n} \ F) = ({n. P n} \ F)" -by (subst Collect_neg_eq, rule Compl_iff) - -lemma (in ultrafilter) Collect_disj: - "({n. P n \ Q n} \ F) = ({n. P n} \ F \ {n. Q n} \ F)" -by (subst Collect_disj_eq, rule Un_iff) - -lemma (in ultrafilter) Collect_all: - "({n. \x. P n x} \ F) = (\X. {n. P n (X n)} \ F)" -apply (rule Not_eq_iff [THEN iffD1]) -apply (simp add: Collect_not [symmetric]) -apply (rule Collect_ex) -done - -subsection {* Maximal filter = Ultrafilter *} - -text {* - A filter F is an ultrafilter iff it is a maximal filter, - i.e. whenever G is a filter and @{term "F \ G"} then @{term "F = G"} -*} -text {* - Lemmas that shows existence of an extension to what was assumed to - be a maximal filter. Will be used to derive contradiction in proof of - property of ultrafilter. -*} - -lemma extend_lemma1: "UNIV \ F \ A \ {X. \f\F. A \ f \ X}" -by blast - -lemma extend_lemma2: "F \ {X. \f\F. A \ f \ X}" -by blast - -lemma (in filter) extend_filter: -assumes A: "- A \ F" -shows "filter {X. \f\F. A \ f \ X}" (is "filter ?X") -proof (rule filter.intro) - show "UNIV \ ?X" by blast -next - show "{} \ ?X" - proof (clarify) - fix f assume f: "f \ F" and Af: "A \ f \ {}" - from Af have fA: "f \ - A" by blast - from f fA have "- A \ F" by (rule subset) - with A show "False" by simp - qed -next - fix u and v - assume u: "u \ ?X" and v: "v \ ?X" - from u obtain f where f: "f \ F" and Af: "A \ f \ u" by blast - from v obtain g where g: "g \ F" and Ag: "A \ g \ v" by blast - from f g have fg: "f \ g \ F" by (rule Int) - from Af Ag have Afg: "A \ (f \ g) \ u \ v" by blast - from fg Afg show "u \ v \ ?X" by blast -next - fix u and v - assume uv: "u \ v" and u: "u \ ?X" - from u obtain f where f: "f \ F" and Afu: "A \ f \ u" by blast - from Afu uv have Afv: "A \ f \ v" by blast - from f Afv have "\f\F. A \ f \ v" by blast - thus "v \ ?X" by simp -qed - -lemma (in filter) max_filter_ultrafilter: -assumes max: "\G. \filter G; F \ G\ \ F = G" -shows "ultrafilter_axioms F" -proof (rule ultrafilter_axioms.intro) - fix A show "A \ F \ - A \ F" - proof (rule disjCI) - let ?X = "{X. \f\F. A \ f \ X}" - assume AF: "- A \ F" - from AF have X: "filter ?X" by (rule extend_filter) - from UNIV have AX: "A \ ?X" by (rule extend_lemma1) - have FX: "F \ ?X" by (rule extend_lemma2) - from X FX have "F = ?X" by (rule max) - with AX show "A \ F" by simp - qed -qed - -lemma (in ultrafilter) max_filter: -assumes G: "filter G" and sub: "F \ G" shows "F = G" -proof - show "F \ G" using sub . - show "G \ F" - proof - fix A assume A: "A \ G" - from G A have "- A \ G" by (rule filter.memD) - with sub have B: "- A \ F" by blast - thus "A \ F" by (rule memI) - qed -qed - -subsection {* Ultrafilter Theorem *} - -text "A local context makes proof of ultrafilter Theorem more modular" -context - fixes frechet :: "'a set set" - and superfrechet :: "'a set set set" - - assumes infinite_UNIV: "infinite (UNIV :: 'a set)" - - defines frechet_def: "frechet \ {A. finite (- A)}" - and superfrechet_def: "superfrechet \ {G. filter G \ frechet \ G}" -begin - -lemma superfrechetI: - "\filter G; frechet \ G\ \ G \ superfrechet" -by (simp add: superfrechet_def) - -lemma superfrechetD1: - "G \ superfrechet \ filter G" -by (simp add: superfrechet_def) - -lemma superfrechetD2: - "G \ superfrechet \ frechet \ G" -by (simp add: superfrechet_def) - -text {* A few properties of free filters *} - -lemma filter_cofinite: -assumes inf: "infinite (UNIV :: 'a set)" -shows "filter {A:: 'a set. finite (- A)}" (is "filter ?F") -proof (rule filter.intro) - show "UNIV \ ?F" by simp -next - show "{} \ ?F" using inf by simp -next - fix u v assume "u \ ?F" and "v \ ?F" - thus "u \ v \ ?F" by simp -next - fix u v assume uv: "u \ v" and u: "u \ ?F" - from uv have vu: "- v \ - u" by simp - from u show "v \ ?F" - by (simp add: finite_subset [OF vu]) -qed - -text {* - We prove: 1. Existence of maximal filter i.e. ultrafilter; - 2. Freeness property i.e ultrafilter is free. - Use a locale to prove various lemmas and then - export main result: The ultrafilter Theorem -*} - -lemma filter_frechet: "filter frechet" -by (unfold frechet_def, rule filter_cofinite [OF infinite_UNIV]) - -lemma frechet_in_superfrechet: "frechet \ superfrechet" -by (rule superfrechetI [OF filter_frechet subset_refl]) - -lemma lemma_mem_chain_filter: - "\c \ chains superfrechet; x \ c\ \ filter x" -by (unfold chains_def superfrechet_def, blast) - - -subsubsection {* Unions of chains of superfrechets *} - -text "In this section we prove that superfrechet is closed -with respect to unions of non-empty chains. We must show - 1) Union of a chain is a filter, - 2) Union of a chain contains frechet. - -Number 2 is trivial, but 1 requires us to prove all the filter rules." - -lemma Union_chain_UNIV: - "\c \ chains superfrechet; c \ {}\ \ UNIV \ \c" -proof - - assume 1: "c \ chains superfrechet" and 2: "c \ {}" - from 2 obtain x where 3: "x \ c" by blast - from 1 3 have "filter x" by (rule lemma_mem_chain_filter) - hence "UNIV \ x" by (rule filter.UNIV) - with 3 show "UNIV \ \c" by blast -qed - -lemma Union_chain_empty: - "c \ chains superfrechet \ {} \ \c" -proof - assume 1: "c \ chains superfrechet" and 2: "{} \ \c" - from 2 obtain x where 3: "x \ c" and 4: "{} \ x" .. - from 1 3 have "filter x" by (rule lemma_mem_chain_filter) - hence "{} \ x" by (rule filter.empty) - with 4 show "False" by simp -qed - -lemma Union_chain_Int: - "\c \ chains superfrechet; u \ \c; v \ \c\ \ u \ v \ \c" -proof - - assume c: "c \ chains superfrechet" - assume "u \ \c" - then obtain x where ux: "u \ x" and xc: "x \ c" .. - assume "v \ \c" - then obtain y where vy: "v \ y" and yc: "y \ c" .. - from c xc yc have "x \ y \ y \ x" using c unfolding chains_def chain_subset_def by auto - with xc yc have xyc: "x \ y \ c" - by (auto simp add: Un_absorb1 Un_absorb2) - with c have fxy: "filter (x \ y)" by (rule lemma_mem_chain_filter) - from ux have uxy: "u \ x \ y" by simp - from vy have vxy: "v \ x \ y" by simp - from fxy uxy vxy have "u \ v \ x \ y" by (rule filter.Int) - with xyc show "u \ v \ \c" .. -qed - -lemma Union_chain_subset: - "\c \ chains superfrechet; u \ \c; u \ v\ \ v \ \c" -proof - - assume c: "c \ chains superfrechet" - and u: "u \ \c" and uv: "u \ v" - from u obtain x where ux: "u \ x" and xc: "x \ c" .. - from c xc have fx: "filter x" by (rule lemma_mem_chain_filter) - from fx ux uv have vx: "v \ x" by (rule filter.subset) - with xc show "v \ \c" .. -qed - -lemma Union_chain_filter: -assumes chain: "c \ chains superfrechet" and nonempty: "c \ {}" -shows "filter (\c)" -proof (rule filter.intro) - show "UNIV \ \c" using chain nonempty by (rule Union_chain_UNIV) -next - show "{} \ \c" using chain by (rule Union_chain_empty) -next - fix u v assume "u \ \c" and "v \ \c" - with chain show "u \ v \ \c" by (rule Union_chain_Int) -next - fix u v assume "u \ \c" and "u \ v" - with chain show "v \ \c" by (rule Union_chain_subset) -qed - -lemma lemma_mem_chain_frechet_subset: - "\c \ chains superfrechet; x \ c\ \ frechet \ x" -by (unfold superfrechet_def chains_def, blast) - -lemma Union_chain_superfrechet: - "\c \ {}; c \ chains superfrechet\ \ \c \ superfrechet" -proof (rule superfrechetI) - assume 1: "c \ chains superfrechet" and 2: "c \ {}" - thus "filter (\c)" by (rule Union_chain_filter) - from 2 obtain x where 3: "x \ c" by blast - from 1 3 have "frechet \ x" by (rule lemma_mem_chain_frechet_subset) - also from 3 have "x \ \c" by blast - finally show "frechet \ \c" . -qed - -subsubsection {* Existence of free ultrafilter *} - -lemma max_cofinite_filter_Ex: - "\U\superfrechet. \G\superfrechet. U \ G \ G = U" -proof (rule Zorn_Lemma2, safe) - fix c assume c: "c \ chains superfrechet" - show "\U\superfrechet. \G\c. G \ U" (is "?U") - proof (cases) - assume "c = {}" - with frechet_in_superfrechet show "?U" by blast - next - assume A: "c \ {}" - from A c have "\c \ superfrechet" - by (rule Union_chain_superfrechet) - thus "?U" by blast - qed -qed - -lemma mem_superfrechet_all_infinite: - "\U \ superfrechet; A \ U\ \ infinite A" -proof - assume U: "U \ superfrechet" and A: "A \ U" and fin: "finite A" - from U have fil: "filter U" and fre: "frechet \ U" - by (simp_all add: superfrechet_def) - from fin have "- A \ frechet" by (simp add: frechet_def) - with fre have cA: "- A \ U" by (rule subsetD) - from fil A cA have "A \ - A \ U" by (rule filter.Int) - with fil show "False" by (simp add: filter.empty) -qed - -text {* There exists a free ultrafilter on any infinite set *} - -lemma freeultrafilter_Ex: - "\U::'a set set. freeultrafilter U" -proof - - from max_cofinite_filter_Ex obtain U - where U: "U \ superfrechet" - and max [rule_format]: "\G\superfrechet. U \ G \ G = U" .. - from U have fil: "filter U" by (rule superfrechetD1) - from U have fre: "frechet \ U" by (rule superfrechetD2) - have ultra: "ultrafilter_axioms U" - proof (rule filter.max_filter_ultrafilter [OF fil]) - fix G assume G: "filter G" and UG: "U \ G" - from fre UG have "frechet \ G" by simp - with G have "G \ superfrechet" by (rule superfrechetI) - from this UG show "U = G" by (rule max[symmetric]) - qed - have free: "freeultrafilter_axioms U" - proof (rule freeultrafilter_axioms.intro) - fix A assume "A \ U" - with U show "infinite A" by (rule mem_superfrechet_all_infinite) - qed - from fil ultra free have "freeultrafilter U" - by (rule freeultrafilter.intro [OF ultrafilter.intro]) - (* FIXME: unfold_locales should use chained facts *) - then show ?thesis .. -qed - -end - -hide_const (open) filter - -end diff -r 4b77fc0b3514 -r 218fcc645d22 src/HOL/NSA/Free_Ultrafilter.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/NSA/Free_Ultrafilter.thy Sun Apr 12 11:33:19 2015 +0200 @@ -0,0 +1,413 @@ +(* Title: HOL/NSA/Free_Ultrafilter.thy + Author: Jacques D. Fleuriot, University of Cambridge + Author: Lawrence C Paulson + Author: Brian Huffman +*) + +section {* Filters and Ultrafilters *} + +theory Free_Ultrafilter +imports "~~/src/HOL/Library/Infinite_Set" +begin + +subsection {* Definitions and basic properties *} + +subsubsection {* Filters *} + +locale filter = + fixes F :: "'a set set" + assumes UNIV [iff]: "UNIV \ F" + assumes empty [iff]: "{} \ F" + assumes Int: "\u \ F; v \ F\ \ u \ v \ F" + assumes subset: "\u \ F; u \ v\ \ v \ F" +begin + +lemma memD: "A \ F \ - A \ F" +proof + assume "A \ F" and "- A \ F" + hence "A \ (- A) \ F" by (rule Int) + thus "False" by simp +qed + +lemma not_memI: "- A \ F \ A \ F" +by (drule memD, simp) + +lemma Int_iff: "(x \ y \ F) = (x \ F \ y \ F)" +by (auto elim: subset intro: Int) + +end + +subsubsection {* Ultrafilters *} + +locale ultrafilter = filter + + assumes ultra: "A \ F \ - A \ F" +begin + +lemma memI: "- A \ F \ A \ F" +using ultra [of A] by simp + +lemma not_memD: "A \ F \ - A \ F" +by (rule memI, simp) + +lemma not_mem_iff: "(A \ F) = (- A \ F)" +by (rule iffI [OF not_memD not_memI]) + +lemma Compl_iff: "(- A \ F) = (A \ F)" +by (rule iffI [OF not_memI not_memD]) + +lemma Un_iff: "(x \ y \ F) = (x \ F \ y \ F)" + apply (rule iffI) + apply (erule contrapos_pp) + apply (simp add: Int_iff not_mem_iff) + apply (auto elim: subset) +done + +end + +subsubsection {* Free Ultrafilters *} + +locale freeultrafilter = ultrafilter + + assumes infinite: "A \ F \ infinite A" +begin + +lemma finite: "finite A \ A \ F" +by (erule contrapos_pn, erule infinite) + +lemma singleton: "{x} \ F" +by (rule finite, simp) + +lemma insert_iff [simp]: "(insert x A \ F) = (A \ F)" +apply (subst insert_is_Un) +apply (subst Un_iff) +apply (simp add: singleton) +done + +lemma filter: "filter F" .. + +lemma ultrafilter: "ultrafilter F" .. + +end + +subsection {* Collect properties *} + +lemma (in filter) Collect_ex: + "({n. \x. P n x} \ F) = (\X. {n. P n (X n)} \ F)" +proof + assume "{n. \x. P n x} \ F" + hence "{n. P n (SOME x. P n x)} \ F" + by (auto elim: someI subset) + thus "\X. {n. P n (X n)} \ F" by fast +next + show "\X. {n. P n (X n)} \ F \ {n. \x. P n x} \ F" + by (auto elim: subset) +qed + +lemma (in filter) Collect_conj: + "({n. P n \ Q n} \ F) = ({n. P n} \ F \ {n. Q n} \ F)" +by (subst Collect_conj_eq, rule Int_iff) + +lemma (in ultrafilter) Collect_not: + "({n. \ P n} \ F) = ({n. P n} \ F)" +by (subst Collect_neg_eq, rule Compl_iff) + +lemma (in ultrafilter) Collect_disj: + "({n. P n \ Q n} \ F) = ({n. P n} \ F \ {n. Q n} \ F)" +by (subst Collect_disj_eq, rule Un_iff) + +lemma (in ultrafilter) Collect_all: + "({n. \x. P n x} \ F) = (\X. {n. P n (X n)} \ F)" +apply (rule Not_eq_iff [THEN iffD1]) +apply (simp add: Collect_not [symmetric]) +apply (rule Collect_ex) +done + +subsection {* Maximal filter = Ultrafilter *} + +text {* + A filter F is an ultrafilter iff it is a maximal filter, + i.e. whenever G is a filter and @{term "F \ G"} then @{term "F = G"} +*} +text {* + Lemmas that shows existence of an extension to what was assumed to + be a maximal filter. Will be used to derive contradiction in proof of + property of ultrafilter. +*} + +lemma extend_lemma1: "UNIV \ F \ A \ {X. \f\F. A \ f \ X}" +by blast + +lemma extend_lemma2: "F \ {X. \f\F. A \ f \ X}" +by blast + +lemma (in filter) extend_filter: +assumes A: "- A \ F" +shows "filter {X. \f\F. A \ f \ X}" (is "filter ?X") +proof (rule filter.intro) + show "UNIV \ ?X" by blast +next + show "{} \ ?X" + proof (clarify) + fix f assume f: "f \ F" and Af: "A \ f \ {}" + from Af have fA: "f \ - A" by blast + from f fA have "- A \ F" by (rule subset) + with A show "False" by simp + qed +next + fix u and v + assume u: "u \ ?X" and v: "v \ ?X" + from u obtain f where f: "f \ F" and Af: "A \ f \ u" by blast + from v obtain g where g: "g \ F" and Ag: "A \ g \ v" by blast + from f g have fg: "f \ g \ F" by (rule Int) + from Af Ag have Afg: "A \ (f \ g) \ u \ v" by blast + from fg Afg show "u \ v \ ?X" by blast +next + fix u and v + assume uv: "u \ v" and u: "u \ ?X" + from u obtain f where f: "f \ F" and Afu: "A \ f \ u" by blast + from Afu uv have Afv: "A \ f \ v" by blast + from f Afv have "\f\F. A \ f \ v" by blast + thus "v \ ?X" by simp +qed + +lemma (in filter) max_filter_ultrafilter: +assumes max: "\G. \filter G; F \ G\ \ F = G" +shows "ultrafilter_axioms F" +proof (rule ultrafilter_axioms.intro) + fix A show "A \ F \ - A \ F" + proof (rule disjCI) + let ?X = "{X. \f\F. A \ f \ X}" + assume AF: "- A \ F" + from AF have X: "filter ?X" by (rule extend_filter) + from UNIV have AX: "A \ ?X" by (rule extend_lemma1) + have FX: "F \ ?X" by (rule extend_lemma2) + from X FX have "F = ?X" by (rule max) + with AX show "A \ F" by simp + qed +qed + +lemma (in ultrafilter) max_filter: +assumes G: "filter G" and sub: "F \ G" shows "F = G" +proof + show "F \ G" using sub . + show "G \ F" + proof + fix A assume A: "A \ G" + from G A have "- A \ G" by (rule filter.memD) + with sub have B: "- A \ F" by blast + thus "A \ F" by (rule memI) + qed +qed + +subsection {* Ultrafilter Theorem *} + +text "A local context makes proof of ultrafilter Theorem more modular" +context + fixes frechet :: "'a set set" + and superfrechet :: "'a set set set" + + assumes infinite_UNIV: "infinite (UNIV :: 'a set)" + + defines frechet_def: "frechet \ {A. finite (- A)}" + and superfrechet_def: "superfrechet \ {G. filter G \ frechet \ G}" +begin + +lemma superfrechetI: + "\filter G; frechet \ G\ \ G \ superfrechet" +by (simp add: superfrechet_def) + +lemma superfrechetD1: + "G \ superfrechet \ filter G" +by (simp add: superfrechet_def) + +lemma superfrechetD2: + "G \ superfrechet \ frechet \ G" +by (simp add: superfrechet_def) + +text {* A few properties of free filters *} + +lemma filter_cofinite: +assumes inf: "infinite (UNIV :: 'a set)" +shows "filter {A:: 'a set. finite (- A)}" (is "filter ?F") +proof (rule filter.intro) + show "UNIV \ ?F" by simp +next + show "{} \ ?F" using inf by simp +next + fix u v assume "u \ ?F" and "v \ ?F" + thus "u \ v \ ?F" by simp +next + fix u v assume uv: "u \ v" and u: "u \ ?F" + from uv have vu: "- v \ - u" by simp + from u show "v \ ?F" + by (simp add: finite_subset [OF vu]) +qed + +text {* + We prove: 1. Existence of maximal filter i.e. ultrafilter; + 2. Freeness property i.e ultrafilter is free. + Use a locale to prove various lemmas and then + export main result: The ultrafilter Theorem +*} + +lemma filter_frechet: "filter frechet" +by (unfold frechet_def, rule filter_cofinite [OF infinite_UNIV]) + +lemma frechet_in_superfrechet: "frechet \ superfrechet" +by (rule superfrechetI [OF filter_frechet subset_refl]) + +lemma lemma_mem_chain_filter: + "\c \ chains superfrechet; x \ c\ \ filter x" +by (unfold chains_def superfrechet_def, blast) + + +subsubsection {* Unions of chains of superfrechets *} + +text "In this section we prove that superfrechet is closed +with respect to unions of non-empty chains. We must show + 1) Union of a chain is a filter, + 2) Union of a chain contains frechet. + +Number 2 is trivial, but 1 requires us to prove all the filter rules." + +lemma Union_chain_UNIV: + "\c \ chains superfrechet; c \ {}\ \ UNIV \ \c" +proof - + assume 1: "c \ chains superfrechet" and 2: "c \ {}" + from 2 obtain x where 3: "x \ c" by blast + from 1 3 have "filter x" by (rule lemma_mem_chain_filter) + hence "UNIV \ x" by (rule filter.UNIV) + with 3 show "UNIV \ \c" by blast +qed + +lemma Union_chain_empty: + "c \ chains superfrechet \ {} \ \c" +proof + assume 1: "c \ chains superfrechet" and 2: "{} \ \c" + from 2 obtain x where 3: "x \ c" and 4: "{} \ x" .. + from 1 3 have "filter x" by (rule lemma_mem_chain_filter) + hence "{} \ x" by (rule filter.empty) + with 4 show "False" by simp +qed + +lemma Union_chain_Int: + "\c \ chains superfrechet; u \ \c; v \ \c\ \ u \ v \ \c" +proof - + assume c: "c \ chains superfrechet" + assume "u \ \c" + then obtain x where ux: "u \ x" and xc: "x \ c" .. + assume "v \ \c" + then obtain y where vy: "v \ y" and yc: "y \ c" .. + from c xc yc have "x \ y \ y \ x" using c unfolding chains_def chain_subset_def by auto + with xc yc have xyc: "x \ y \ c" + by (auto simp add: Un_absorb1 Un_absorb2) + with c have fxy: "filter (x \ y)" by (rule lemma_mem_chain_filter) + from ux have uxy: "u \ x \ y" by simp + from vy have vxy: "v \ x \ y" by simp + from fxy uxy vxy have "u \ v \ x \ y" by (rule filter.Int) + with xyc show "u \ v \ \c" .. +qed + +lemma Union_chain_subset: + "\c \ chains superfrechet; u \ \c; u \ v\ \ v \ \c" +proof - + assume c: "c \ chains superfrechet" + and u: "u \ \c" and uv: "u \ v" + from u obtain x where ux: "u \ x" and xc: "x \ c" .. + from c xc have fx: "filter x" by (rule lemma_mem_chain_filter) + from fx ux uv have vx: "v \ x" by (rule filter.subset) + with xc show "v \ \c" .. +qed + +lemma Union_chain_filter: +assumes chain: "c \ chains superfrechet" and nonempty: "c \ {}" +shows "filter (\c)" +proof (rule filter.intro) + show "UNIV \ \c" using chain nonempty by (rule Union_chain_UNIV) +next + show "{} \ \c" using chain by (rule Union_chain_empty) +next + fix u v assume "u \ \c" and "v \ \c" + with chain show "u \ v \ \c" by (rule Union_chain_Int) +next + fix u v assume "u \ \c" and "u \ v" + with chain show "v \ \c" by (rule Union_chain_subset) +qed + +lemma lemma_mem_chain_frechet_subset: + "\c \ chains superfrechet; x \ c\ \ frechet \ x" +by (unfold superfrechet_def chains_def, blast) + +lemma Union_chain_superfrechet: + "\c \ {}; c \ chains superfrechet\ \ \c \ superfrechet" +proof (rule superfrechetI) + assume 1: "c \ chains superfrechet" and 2: "c \ {}" + thus "filter (\c)" by (rule Union_chain_filter) + from 2 obtain x where 3: "x \ c" by blast + from 1 3 have "frechet \ x" by (rule lemma_mem_chain_frechet_subset) + also from 3 have "x \ \c" by blast + finally show "frechet \ \c" . +qed + +subsubsection {* Existence of free ultrafilter *} + +lemma max_cofinite_filter_Ex: + "\U\superfrechet. \G\superfrechet. U \ G \ G = U" +proof (rule Zorn_Lemma2, safe) + fix c assume c: "c \ chains superfrechet" + show "\U\superfrechet. \G\c. G \ U" (is "?U") + proof (cases) + assume "c = {}" + with frechet_in_superfrechet show "?U" by blast + next + assume A: "c \ {}" + from A c have "\c \ superfrechet" + by (rule Union_chain_superfrechet) + thus "?U" by blast + qed +qed + +lemma mem_superfrechet_all_infinite: + "\U \ superfrechet; A \ U\ \ infinite A" +proof + assume U: "U \ superfrechet" and A: "A \ U" and fin: "finite A" + from U have fil: "filter U" and fre: "frechet \ U" + by (simp_all add: superfrechet_def) + from fin have "- A \ frechet" by (simp add: frechet_def) + with fre have cA: "- A \ U" by (rule subsetD) + from fil A cA have "A \ - A \ U" by (rule filter.Int) + with fil show "False" by (simp add: filter.empty) +qed + +text {* There exists a free ultrafilter on any infinite set *} + +lemma freeultrafilter_Ex: + "\U::'a set set. freeultrafilter U" +proof - + from max_cofinite_filter_Ex obtain U + where U: "U \ superfrechet" + and max [rule_format]: "\G\superfrechet. U \ G \ G = U" .. + from U have fil: "filter U" by (rule superfrechetD1) + from U have fre: "frechet \ U" by (rule superfrechetD2) + have ultra: "ultrafilter_axioms U" + proof (rule filter.max_filter_ultrafilter [OF fil]) + fix G assume G: "filter G" and UG: "U \ G" + from fre UG have "frechet \ G" by simp + with G have "G \ superfrechet" by (rule superfrechetI) + from this UG show "U = G" by (rule max[symmetric]) + qed + have free: "freeultrafilter_axioms U" + proof (rule freeultrafilter_axioms.intro) + fix A assume "A \ U" + with U show "infinite A" by (rule mem_superfrechet_all_infinite) + qed + from fil ultra free have "freeultrafilter U" + by (rule freeultrafilter.intro [OF ultrafilter.intro]) + (* FIXME: unfold_locales should use chained facts *) + then show ?thesis .. +qed + +end + +hide_const (open) filter + +end diff -r 4b77fc0b3514 -r 218fcc645d22 src/HOL/NSA/StarDef.thy --- a/src/HOL/NSA/StarDef.thy Sun Apr 12 16:04:53 2015 +0200 +++ b/src/HOL/NSA/StarDef.thy Sun Apr 12 11:33:19 2015 +0200 @@ -5,7 +5,7 @@ section {* Construction of Star Types Using Ultrafilters *} theory StarDef -imports Filter +imports Free_Ultrafilter begin subsection {* A Free Ultrafilter over the Naturals *} diff -r 4b77fc0b3514 -r 218fcc645d22 src/HOL/Topological_Spaces.thy --- a/src/HOL/Topological_Spaces.thy Sun Apr 12 16:04:53 2015 +0200 +++ b/src/HOL/Topological_Spaces.thy Sun Apr 12 11:33:19 2015 +0200 @@ -322,557 +322,6 @@ by auto qed -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" - -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 eventually_mono: - "(\x. P x \ Q x) \ eventually P F \ eventually Q F" - unfolding eventually_def - by (rule 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_Ball_finite: - assumes "finite A" and "\y\A. eventually (\x. P x y) net" - shows "eventually (\x. \y\A. P x y) net" -using assms by (induct set: finite, simp, simp add: eventually_conj) - -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_mp: - assumes "eventually (\x. P x \ Q x) F" - assumes "eventually (\x. P x) F" - shows "eventually (\x. Q x) F" -proof (rule eventually_mono) - show "\x. (P x \ Q x) \ P x \ Q x" by simp - show "eventually (\x. (P x \ Q x) \ P x) F" - using assms by (rule eventually_conj) -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_elim1: - assumes "eventually (\i. P i) F" - assumes "\i. P i \ Q i" - shows "eventually (\i. Q i) F" - using assms by (auto 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 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_elim1) - then show ?thesis by (auto elim: eventually_elim2) -qed - -ML {* - fun eventually_elim_tac ctxt facts = SUBGOAL_CASES (fn (goal, i) => - let - val mp_thms = facts RL @{thms eventually_rev_mp} - val raw_elim_thm = - (@{thm allI} RS @{thm always_eventually}) - |> fold (fn thm1 => fn thm2 => thm2 RS thm1) mp_thms - |> fold (fn _ => fn thm => @{thm impI} RS thm) facts - val cases_prop = Thm.prop_of (raw_elim_thm RS Goal.init (Thm.cterm_of ctxt goal)) - val cases = Rule_Cases.make_common ctxt cases_prop [(("elim", []), [])] - in - CASES cases (rtac raw_elim_thm i) - end) -*} - -method_setup eventually_elim = {* - Scan.succeed (fn ctxt => METHOD_CASES (HEADGOAL o eventually_elim_tac ctxt)) -*} "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, rule is_filter.intro) - apply (fast intro: eventually_True) - apply clarify - apply (intro exI conjI) - apply (erule (1) eventually_conj) - apply (erule (1) eventually_conj) - apply simp - apply auto - done - -lemma eventually_Sup: - "eventually P (Sup S) \ (\F\S. eventually P F)" - unfolding Sup_filter_def - apply (rule eventually_Abs_filter, rule 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 elim!: eventually_mono intro: 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 - -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) - -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 eventually_const: "\ trivial_limit net \ eventually (\x. P) net \ P" - by (cases P) (simp_all 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_elim1) - qed } - note eventually_F = this - - 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 (INF b:B. F b) \ (\X\B. finite X \ eventually P (INF b:X. F b))" - unfolding INF_def[of B] eventually_Inf[of P "F`B"] - by (metis Inf_image_eq 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 \ (INF b:X. F b) \ bot) \ (INF 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 (INF b:B. F b) \ (\b\B. eventually P (F b))" - unfolding INF_def by (subst eventually_Inf_base) 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) - apply (rule 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_sup: "filtermap f (sup F1 F2) = sup (filtermap f F1) (filtermap f F2)" - by (auto simp: filter_eq_iff eventually_filtermap eventually_sup) - -lemma filtermap_inf: "filtermap f (inf F1 F2) \ inf (filtermap f F1) (filtermap f F2)" - by (auto simp: le_filter_def eventually_filtermap eventually_inf) - -lemma filtermap_INF: "filtermap f (INF b:B. F b) \ (INF b:B. filtermap f (F b))" -proof - - { fix X :: "'c set" assume "finite X" - then have "filtermap f (INFIMUM X F) \ (INF b:X. filtermap f (F b))" - proof induct - case (insert x X) - have "filtermap f (INF a:insert x X. F a) \ inf (filtermap f (F x)) (filtermap f (INF a:X. F a))" - by (rule order_trans[OF _ filtermap_inf]) simp - also have "\ \ inf (filtermap f (F x)) (INF a:X. filtermap f (F a))" - by (intro inf_mono insert order_refl) - finally show ?case - by simp - qed simp } - then show ?thesis - unfolding le_filter_def eventually_filtermap - by (subst (1 2) eventually_INF) auto -qed -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_elim1) - -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 - apply safe - apply (erule_tac x="\x. x \ A" in allE) - apply (auto elim: eventually_elim1) - done - -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]: "(SUP i : I. principal (A i)) = principal (\i\I. A i)" - unfolding filter_eq_iff eventually_Sup SUP_def by (auto simp: eventually_principal) - -lemma INF_principal_finite: "finite X \ (INF 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 - -subsubsection {* Order filters *} - -definition at_top :: "('a::order) filter" - where "at_top = (INF k. principal {k ..})" - -lemma at_top_sub: "at_top = (INF 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_ge_at_top: - "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 (INF 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 "(INF 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_gt_at_top: - "eventually (\x. (c::_::unbounded_dense_linorder) < x) at_top" - unfolding eventually_at_top_dense by auto - -definition at_bot :: "('a::order) filter" - where "at_bot = (INF k. principal {.. k})" - -lemma at_bot_sub: "at_bot = (INF 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_le_at_bot: - "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}. \n (\N::'a. \nx. x < (c::_::unbounded_dense_linorder)) at_bot" - unfolding eventually_at_bot_dense by auto - -lemma trivial_limit_at_bot_linorder: "\ 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: "\ 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: - assumes "\x. c \ x \ P x" - shows "eventually P sequentially" -using assms by (auto simp: eventually_sequentially) - -lemma eventually_sequentially_seg: - "eventually (\n. P (n + k)) sequentially \ eventually P sequentially" - unfolding eventually_sequentially - apply safe - apply (rule_tac x="N + k" in exI) - apply rule - apply (erule_tac x="n - k" in allE) - apply auto [] - apply (rule_tac x=N in exI) - apply auto [] - done - subsubsection {* Topological filters *} definition (in topological_space) nhds :: "'a \ 'a filter" @@ -1005,104 +454,6 @@ "eventually P (at (x::'a::linorder_topology)) \ eventually P (at_left x) \ eventually P (at_right x)" by (subst at_eq_sup_left_right) (simp add: eventually_sup) -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_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'" - apply (rule filterlim_cong[OF refl refl eq, THEN iffD1]) - apply (rule filterlim_mono[OF _ ord]) - apply fact - done - -lemma filtermap_mono_strong: "inj f \ filtermap f F \ filtermap f G \ F \ G" - apply (auto 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 filtermap_eq_strong: "inj f \ filtermap f F = filtermap f G \ F = G" - by (simp add: filtermap_mono_strong eq_iff) - -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_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 :> (INF 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 (INF i:I. F i). f x :> (INF j:J. G j)" - unfolding filterlim_def by (rule order_trans[OF filtermap_INF INF_mono]) - -lemma filterlim_base: - "(\m x. m \ J \ i m \ I) \ (\m x. m \ J \ x \ F (i m) \ f x \ G m) \ - LIM x (INF i:I. principal (F i)). f x :> (INF 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 (INF i:I. principal (F i)). f x :> INF 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 eventually_sequentially_Suc: "eventually (\i. P (Suc i)) sequentially \ eventually P sequentially" - unfolding eventually_sequentially by (metis Suc_le_D Suc_le_mono le_Suc_eq) - -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) (metis le_Suc_eq) - subsubsection {* Tendsto *} abbreviation (in topological_space) @@ -1296,92 +647,6 @@ lemma Lim_ident_at: "\ trivial_limit (at x within s) \ Lim (at x within s) (\x. x) = x" by (rule tendsto_Lim[OF _ tendsto_ident_at]) auto -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_elim1) - -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_elim1[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_elim1) - -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" - apply (rule eventually_mono[rotated]) - using 1 by 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[rotated], 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_elim1) -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_at_bot_at_right: fixes f :: "'a::linorder_topology \ 'b::linorder" assumes mono: "\x y. Q x \ Q y \ x \ y \ f x \ f y" @@ -2936,220 +2201,4 @@ qed qed (intro cSUP_least `antimono f`[THEN antimonoD] cInf_lower S) -subsection {* Setup @{typ "'a filter"} for lifting and transfer *} - -context begin interpretation lifting_syntax . - -definition rel_filter :: "('a \ 'b \ bool) \ 'a filter \ 'b filter \ bool" -where "rel_filter R F G = ((R ===> op =) ===> op =) (Rep_filter F) (Rep_filter G)" - -lemma rel_filter_eventually: - "rel_filter R F G \ - ((R ===> op =) ===> op =) (\P. eventually P F) (\P. eventually P G)" -by(simp add: rel_filter_def eventually_def) - -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 . - -lemma Quotient_filter [quot_map]: - assumes Q: "Quotient R Abs Rep T" - shows "Quotient (rel_filter R) (filtermap Abs) (filtermap Rep) (rel_filter T)" -unfolding Quotient_alt_def -proof(intro conjI strip) - from Q have *: "\x y. T x y \ Abs x = y" - unfolding Quotient_alt_def by blast - - fix F G - assume "rel_filter T F G" - thus "filtermap Abs F = G" unfolding filter_eq_iff - by(auto simp add: eventually_filtermap rel_filter_eventually * rel_funI del: iffI elim!: rel_funD) -next - from Q have *: "\x. T (Rep x) x" unfolding Quotient_alt_def by blast - - fix F - show "rel_filter T (filtermap Rep F) F" - by(auto elim: rel_funD intro: * intro!: ext arg_cong[where f="\P. eventually P F"] rel_funI - del: iffI simp add: eventually_filtermap rel_filter_eventually) -qed(auto simp add: map_fun_def o_def eventually_filtermap filter_eq_iff fun_eq_iff rel_filter_eventually - fun_quotient[OF fun_quotient[OF Q identity_quotient] identity_quotient, unfolded Quotient_alt_def]) - -lemma eventually_parametric [transfer_rule]: - "((A ===> op =) ===> rel_filter A ===> op =) eventually eventually" -by(simp add: rel_fun_def rel_filter_eventually) - -lemma rel_filter_eq [relator_eq]: "rel_filter op = = op =" -by(auto simp add: rel_filter_eventually rel_fun_eq fun_eq_iff filter_eq_iff) - -lemma rel_filter_mono [relator_mono]: - "A \ B \ rel_filter A \ rel_filter B" -unfolding rel_filter_eventually[abs_def] -by(rule le_funI)+(intro fun_mono fun_mono[THEN le_funD, THEN le_funD] order.refl) - -lemma rel_filter_conversep [simp]: "rel_filter A\\ = (rel_filter A)\\" -by(auto simp add: rel_filter_eventually fun_eq_iff rel_fun_def) - -lemma is_filter_parametric_aux: - assumes "is_filter F" - assumes [transfer_rule]: "bi_total A" "bi_unique A" - and [transfer_rule]: "((A ===> op =) ===> op =) F G" - shows "is_filter G" -proof - - interpret is_filter F by fact - show ?thesis - proof - have "F (\_. True) = G (\x. True)" by transfer_prover - thus "G (\x. True)" by(simp add: True) - next - fix P' Q' - assume "G P'" "G Q'" - moreover - from bi_total_fun[OF `bi_unique A` bi_total_eq, unfolded bi_total_def] - obtain P Q where [transfer_rule]: "(A ===> op =) P P'" "(A ===> op =) Q Q'" by blast - have "F P = G P'" "F Q = G Q'" by transfer_prover+ - ultimately have "F (\x. P x \ Q x)" by(simp add: conj) - moreover have "F (\x. P x \ Q x) = G (\x. P' x \ Q' x)" by transfer_prover - ultimately show "G (\x. P' x \ Q' x)" by simp - next - fix P' Q' - assume "\x. P' x \ Q' x" "G P'" - moreover - from bi_total_fun[OF `bi_unique A` bi_total_eq, unfolded bi_total_def] - obtain P Q where [transfer_rule]: "(A ===> op =) P P'" "(A ===> op =) Q Q'" by blast - have "F P = G P'" by transfer_prover - moreover have "(\x. P x \ Q x) \ (\x. P' x \ Q' x)" by transfer_prover - ultimately have "F Q" by(simp add: mono) - moreover have "F Q = G Q'" by transfer_prover - ultimately show "G Q'" by simp - qed -qed - -lemma is_filter_parametric [transfer_rule]: - "\ bi_total A; bi_unique A \ - \ (((A ===> op =) ===> op =) ===> op =) is_filter is_filter" -apply(rule rel_funI) -apply(rule iffI) - apply(erule (3) is_filter_parametric_aux) -apply(erule is_filter_parametric_aux[where A="conversep A"]) -apply(auto simp add: rel_fun_def) -done - -lemma left_total_rel_filter [transfer_rule]: - assumes [transfer_rule]: "bi_total A" "bi_unique A" - shows "left_total (rel_filter A)" -proof(rule left_totalI) - fix F :: "'a filter" - from bi_total_fun[OF bi_unique_fun[OF `bi_total A` bi_unique_eq] bi_total_eq] - obtain G where [transfer_rule]: "((A ===> op =) ===> op =) (\P. eventually P F) G" - unfolding bi_total_def by blast - moreover have "is_filter (\P. eventually P F) \ is_filter G" by transfer_prover - hence "is_filter G" by(simp add: eventually_def is_filter_Rep_filter) - ultimately have "rel_filter A F (Abs_filter G)" - by(simp add: rel_filter_eventually eventually_Abs_filter) - thus "\G. rel_filter A F G" .. -qed - -lemma right_total_rel_filter [transfer_rule]: - "\ bi_total A; bi_unique A \ \ right_total (rel_filter A)" -using left_total_rel_filter[of "A\\"] by simp - -lemma bi_total_rel_filter [transfer_rule]: - assumes "bi_total A" "bi_unique A" - shows "bi_total (rel_filter A)" -unfolding bi_total_alt_def using assms -by(simp add: left_total_rel_filter right_total_rel_filter) - -lemma left_unique_rel_filter [transfer_rule]: - assumes "left_unique A" - shows "left_unique (rel_filter A)" -proof(rule left_uniqueI) - fix F F' G - assume [transfer_rule]: "rel_filter A F G" "rel_filter A F' G" - show "F = F'" - unfolding filter_eq_iff - proof - fix P :: "'a \ bool" - obtain P' where [transfer_rule]: "(A ===> op =) P P'" - using left_total_fun[OF assms left_total_eq] unfolding left_total_def by blast - have "eventually P F = eventually P' G" - and "eventually P F' = eventually P' G" by transfer_prover+ - thus "eventually P F = eventually P F'" by simp - qed -qed - -lemma right_unique_rel_filter [transfer_rule]: - "right_unique A \ right_unique (rel_filter A)" -using left_unique_rel_filter[of "A\\"] by simp - -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 top_filter_parametric [transfer_rule]: - "bi_total A \ (rel_filter A) top top" -by(simp add: rel_filter_eventually All_transfer) - -lemma bot_filter_parametric [transfer_rule]: "(rel_filter A) bot bot" -by(simp add: rel_filter_eventually rel_fun_def) - -lemma sup_filter_parametric [transfer_rule]: - "(rel_filter A ===> rel_filter A ===> rel_filter A) sup sup" -by(fastforce simp add: rel_filter_eventually[abs_def] eventually_sup dest: rel_funD) - -lemma Sup_filter_parametric [transfer_rule]: - "(rel_set (rel_filter A) ===> rel_filter A) Sup Sup" -proof(rule rel_funI) - fix S T - assume [transfer_rule]: "rel_set (rel_filter A) S T" - show "rel_filter A (Sup S) (Sup T)" - by(simp add: rel_filter_eventually eventually_Sup) transfer_prover -qed - -lemma principal_parametric [transfer_rule]: - "(rel_set A ===> rel_filter A) principal principal" -proof(rule rel_funI) - fix S S' - assume [transfer_rule]: "rel_set A S S'" - show "rel_filter A (principal S) (principal S')" - by(simp add: rel_filter_eventually eventually_principal) transfer_prover -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 ===> op =) op \ op \" -unfolding le_filter_def[abs_def] by transfer_prover - -lemma less_filter_parametric [transfer_rule]: - "(rel_filter A ===> rel_filter A ===> op =) op < op <" -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 - -end diff -r 4b77fc0b3514 -r 218fcc645d22 src/HOL/Transcendental.thy --- a/src/HOL/Transcendental.thy Sun Apr 12 16:04:53 2015 +0200 +++ b/src/HOL/Transcendental.thy Sun Apr 12 11:33:19 2015 +0200 @@ -2321,7 +2321,7 @@ shows "\(f ---> a) F; (g ---> b) F; a \ 0\ \ ((\x. f x powr g x) ---> a powr b) F" apply (simp add: powr_def) apply (simp add: tendsto_def) - apply (simp add: Topological_Spaces.eventually_conj_iff ) + apply (simp add: eventually_conj_iff ) apply safe apply (case_tac "0 \ S") apply (auto simp: )