diff --git a/src/HOL/Algebra/Divisibility.thy b/src/HOL/Algebra/Divisibility.thy --- a/src/HOL/Algebra/Divisibility.thy +++ b/src/HOL/Algebra/Divisibility.thy @@ -1,2939 +1,2939 @@ (* Title: HOL/Algebra/Divisibility.thy Author: Clemens Ballarin Author: Stephan Hohe *) section \Divisibility in monoids and rings\ theory Divisibility imports "HOL-Combinatorics.List_Permutation" Coset Group begin section \Factorial Monoids\ subsection \Monoids with Cancellation Law\ locale monoid_cancel = monoid + assumes l_cancel: "\c \ a = c \ b; a \ carrier G; b \ carrier G; c \ carrier G\ \ a = b" and r_cancel: "\a \ c = b \ c; a \ carrier G; b \ carrier G; c \ carrier G\ \ a = b" lemma (in monoid) monoid_cancelI: assumes l_cancel: "\a b c. \c \ a = c \ b; a \ carrier G; b \ carrier G; c \ carrier G\ \ a = b" and r_cancel: "\a b c. \a \ c = b \ c; a \ carrier G; b \ carrier G; c \ carrier G\ \ a = b" shows "monoid_cancel G" by standard fact+ lemma (in monoid_cancel) is_monoid_cancel: "monoid_cancel G" .. sublocale group \ monoid_cancel by standard simp_all locale comm_monoid_cancel = monoid_cancel + comm_monoid lemma comm_monoid_cancelI: fixes G (structure) assumes "comm_monoid G" assumes cancel: "\a b c. \a \ c = b \ c; a \ carrier G; b \ carrier G; c \ carrier G\ \ a = b" shows "comm_monoid_cancel G" proof - interpret comm_monoid G by fact show "comm_monoid_cancel G" by unfold_locales (metis assms(2) m_ac(2))+ qed lemma (in comm_monoid_cancel) is_comm_monoid_cancel: "comm_monoid_cancel G" by intro_locales sublocale comm_group \ comm_monoid_cancel .. subsection \Products of Units in Monoids\ lemma (in monoid) prod_unit_l: assumes abunit[simp]: "a \ b \ Units G" and aunit[simp]: "a \ Units G" and carr[simp]: "a \ carrier G" "b \ carrier G" shows "b \ Units G" proof - have c: "inv (a \ b) \ a \ carrier G" by simp have "(inv (a \ b) \ a) \ b = inv (a \ b) \ (a \ b)" by (simp add: m_assoc) also have "\ = \" by simp finally have li: "(inv (a \ b) \ a) \ b = \" . have "\ = inv a \ a" by (simp add: Units_l_inv[symmetric]) also have "\ = inv a \ \ \ a" by simp also have "\ = inv a \ ((a \ b) \ inv (a \ b)) \ a" by (simp add: Units_r_inv[OF abunit, symmetric] del: Units_r_inv) also have "\ = ((inv a \ a) \ b) \ inv (a \ b) \ a" by (simp add: m_assoc del: Units_l_inv) also have "\ = b \ inv (a \ b) \ a" by simp also have "\ = b \ (inv (a \ b) \ a)" by (simp add: m_assoc) finally have ri: "b \ (inv (a \ b) \ a) = \ " by simp from c li ri show "b \ Units G" by (auto simp: Units_def) qed lemma (in monoid) prod_unit_r: assumes abunit[simp]: "a \ b \ Units G" and bunit[simp]: "b \ Units G" and carr[simp]: "a \ carrier G" "b \ carrier G" shows "a \ Units G" proof - have c: "b \ inv (a \ b) \ carrier G" by simp have "a \ (b \ inv (a \ b)) = (a \ b) \ inv (a \ b)" by (simp add: m_assoc del: Units_r_inv) also have "\ = \" by simp finally have li: "a \ (b \ inv (a \ b)) = \" . have "\ = b \ inv b" by (simp add: Units_r_inv[symmetric]) also have "\ = b \ \ \ inv b" by simp also have "\ = b \ (inv (a \ b) \ (a \ b)) \ inv b" by (simp add: Units_l_inv[OF abunit, symmetric] del: Units_l_inv) also have "\ = (b \ inv (a \ b) \ a) \ (b \ inv b)" by (simp add: m_assoc del: Units_l_inv) also have "\ = b \ inv (a \ b) \ a" by simp finally have ri: "(b \ inv (a \ b)) \ a = \ " by simp from c li ri show "a \ Units G" by (auto simp: Units_def) qed lemma (in comm_monoid) unit_factor: assumes abunit: "a \ b \ Units G" and [simp]: "a \ carrier G" "b \ carrier G" shows "a \ Units G" using abunit[simplified Units_def] proof clarsimp fix i assume [simp]: "i \ carrier G" have carr': "b \ i \ carrier G" by simp have "(b \ i) \ a = (i \ b) \ a" by (simp add: m_comm) also have "\ = i \ (b \ a)" by (simp add: m_assoc) also have "\ = i \ (a \ b)" by (simp add: m_comm) also assume "i \ (a \ b) = \" finally have li': "(b \ i) \ a = \" . have "a \ (b \ i) = a \ b \ i" by (simp add: m_assoc) also assume "a \ b \ i = \" finally have ri': "a \ (b \ i) = \" . from carr' li' ri' show "a \ Units G" by (simp add: Units_def, fast) qed subsection \Divisibility and Association\ subsubsection \Function definitions\ definition factor :: "[_, 'a, 'a] \ bool" (infix "divides\" 65) where "a divides\<^bsub>G\<^esub> b \ (\c\carrier G. b = a \\<^bsub>G\<^esub> c)" definition associated :: "[_, 'a, 'a] \ bool" (infix "\\" 55) where "a \\<^bsub>G\<^esub> b \ a divides\<^bsub>G\<^esub> b \ b divides\<^bsub>G\<^esub> a" abbreviation "division_rel G \ \carrier = carrier G, eq = (\\<^bsub>G\<^esub>), le = (divides\<^bsub>G\<^esub>)\" definition properfactor :: "[_, 'a, 'a] \ bool" where "properfactor G a b \ a divides\<^bsub>G\<^esub> b \ \(b divides\<^bsub>G\<^esub> a)" definition irreducible :: "[_, 'a] \ bool" where "irreducible G a \ a \ Units G \ (\b\carrier G. properfactor G b a \ b \ Units G)" definition prime :: "[_, 'a] \ bool" where "prime G p \ p \ Units G \ (\a\carrier G. \b\carrier G. p divides\<^bsub>G\<^esub> (a \\<^bsub>G\<^esub> b) \ p divides\<^bsub>G\<^esub> a \ p divides\<^bsub>G\<^esub> b)" subsubsection \Divisibility\ lemma dividesI: fixes G (structure) assumes carr: "c \ carrier G" and p: "b = a \ c" shows "a divides b" unfolding factor_def using assms by fast lemma dividesI' [intro]: fixes G (structure) assumes p: "b = a \ c" and carr: "c \ carrier G" shows "a divides b" using assms by (fast intro: dividesI) lemma dividesD: fixes G (structure) assumes "a divides b" shows "\c\carrier G. b = a \ c" using assms unfolding factor_def by fast lemma dividesE [elim]: fixes G (structure) assumes d: "a divides b" and elim: "\c. \b = a \ c; c \ carrier G\ \ P" shows "P" proof - from dividesD[OF d] obtain c where "c \ carrier G" and "b = a \ c" by auto then show P by (elim elim) qed lemma (in monoid) divides_refl[simp, intro!]: assumes carr: "a \ carrier G" shows "a divides a" by (intro dividesI[of "\"]) (simp_all add: carr) lemma (in monoid) divides_trans [trans]: assumes dvds: "a divides b" "b divides c" and acarr: "a \ carrier G" shows "a divides c" using dvds[THEN dividesD] by (blast intro: dividesI m_assoc acarr) lemma (in monoid) divides_mult_lI [intro]: assumes "a divides b" "a \ carrier G" "c \ carrier G" shows "(c \ a) divides (c \ b)" by (metis assms factor_def m_assoc) lemma (in monoid_cancel) divides_mult_l [simp]: assumes carr: "a \ carrier G" "b \ carrier G" "c \ carrier G" shows "(c \ a) divides (c \ b) = a divides b" proof show "c \ a divides c \ b \ a divides b" using carr monoid.m_assoc monoid_axioms monoid_cancel.l_cancel monoid_cancel_axioms by fastforce show "a divides b \ c \ a divides c \ b" using carr(1) carr(3) by blast qed lemma (in comm_monoid) divides_mult_rI [intro]: assumes ab: "a divides b" and carr: "a \ carrier G" "b \ carrier G" "c \ carrier G" shows "(a \ c) divides (b \ c)" using carr ab by (metis divides_mult_lI m_comm) lemma (in comm_monoid_cancel) divides_mult_r [simp]: assumes carr: "a \ carrier G" "b \ carrier G" "c \ carrier G" shows "(a \ c) divides (b \ c) = a divides b" using carr by (simp add: m_comm[of a c] m_comm[of b c]) lemma (in monoid) divides_prod_r: assumes ab: "a divides b" and carr: "a \ carrier G" "c \ carrier G" shows "a divides (b \ c)" using ab carr by (fast intro: m_assoc) lemma (in comm_monoid) divides_prod_l: assumes "a \ carrier G" "b \ carrier G" "c \ carrier G" "a divides b" shows "a divides (c \ b)" using assms by (simp add: divides_prod_r m_comm) lemma (in monoid) unit_divides: assumes uunit: "u \ Units G" and acarr: "a \ carrier G" shows "u divides a" proof (intro dividesI[of "(inv u) \ a"], fast intro: uunit acarr) from uunit acarr have xcarr: "inv u \ a \ carrier G" by fast from uunit acarr have "u \ (inv u \ a) = (u \ inv u) \ a" by (fast intro: m_assoc[symmetric]) also have "\ = \ \ a" by (simp add: Units_r_inv[OF uunit]) also from acarr have "\ = a" by simp finally show "a = u \ (inv u \ a)" .. qed lemma (in comm_monoid) divides_unit: assumes udvd: "a divides u" and carr: "a \ carrier G" "u \ Units G" shows "a \ Units G" using udvd carr by (blast intro: unit_factor) lemma (in comm_monoid) Unit_eq_dividesone: assumes ucarr: "u \ carrier G" shows "u \ Units G = u divides \" using ucarr by (fast dest: divides_unit intro: unit_divides) subsubsection \Association\ lemma associatedI: fixes G (structure) assumes "a divides b" "b divides a" shows "a \ b" using assms by (simp add: associated_def) lemma (in monoid) associatedI2: assumes uunit[simp]: "u \ Units G" and a: "a = b \ u" and bcarr: "b \ carrier G" shows "a \ b" using uunit bcarr unfolding a apply (intro associatedI) apply (metis Units_closed divides_mult_lI one_closed r_one unit_divides) by blast lemma (in monoid) associatedI2': assumes "a = b \ u" and "u \ Units G" and "b \ carrier G" shows "a \ b" using assms by (intro associatedI2) lemma associatedD: fixes G (structure) assumes "a \ b" shows "a divides b" using assms by (simp add: associated_def) lemma (in monoid_cancel) associatedD2: assumes assoc: "a \ b" and carr: "a \ carrier G" "b \ carrier G" shows "\u\Units G. a = b \ u" using assoc unfolding associated_def proof clarify assume "b divides a" then obtain u where ucarr: "u \ carrier G" and a: "a = b \ u" by (rule dividesE) assume "a divides b" then obtain u' where u'carr: "u' \ carrier G" and b: "b = a \ u'" by (rule dividesE) note carr = carr ucarr u'carr from carr have "a \ \ = a" by simp also have "\ = b \ u" by (simp add: a) also have "\ = a \ u' \ u" by (simp add: b) also from carr have "\ = a \ (u' \ u)" by (simp add: m_assoc) finally have "a \ \ = a \ (u' \ u)" . with carr have u1: "\ = u' \ u" by (fast dest: l_cancel) from carr have "b \ \ = b" by simp also have "\ = a \ u'" by (simp add: b) also have "\ = b \ u \ u'" by (simp add: a) also from carr have "\ = b \ (u \ u')" by (simp add: m_assoc) finally have "b \ \ = b \ (u \ u')" . with carr have u2: "\ = u \ u'" by (fast dest: l_cancel) from u'carr u1[symmetric] u2[symmetric] have "\u'\carrier G. u' \ u = \ \ u \ u' = \" by fast then have "u \ Units G" by (simp add: Units_def ucarr) with ucarr a show "\u\Units G. a = b \ u" by fast qed lemma associatedE: fixes G (structure) assumes assoc: "a \ b" and e: "\a divides b; b divides a\ \ P" shows "P" proof - from assoc have "a divides b" "b divides a" by (simp_all add: associated_def) then show P by (elim e) qed lemma (in monoid_cancel) associatedE2: assumes assoc: "a \ b" and e: "\u. \a = b \ u; u \ Units G\ \ P" and carr: "a \ carrier G" "b \ carrier G" shows "P" proof - from assoc and carr have "\u\Units G. a = b \ u" by (rule associatedD2) then obtain u where "u \ Units G" "a = b \ u" by auto then show P by (elim e) qed lemma (in monoid) associated_refl [simp, intro!]: assumes "a \ carrier G" shows "a \ a" using assms by (fast intro: associatedI) lemma (in monoid) associated_sym [sym]: assumes "a \ b" shows "b \ a" using assms by (iprover intro: associatedI elim: associatedE) lemma (in monoid) associated_trans [trans]: assumes "a \ b" "b \ c" and "a \ carrier G" "c \ carrier G" shows "a \ c" using assms by (iprover intro: associatedI divides_trans elim: associatedE) lemma (in monoid) division_equiv [intro, simp]: "equivalence (division_rel G)" apply unfold_locales apply simp_all apply (metis associated_def) apply (iprover intro: associated_trans) done subsubsection \Division and associativity\ lemmas divides_antisym = associatedI lemma (in monoid) divides_cong_l [trans]: assumes "x \ x'" "x' divides y" "x \ carrier G" shows "x divides y" by (meson assms associatedD divides_trans) lemma (in monoid) divides_cong_r [trans]: assumes "x divides y" "y \ y'" "x \ carrier G" shows "x divides y'" by (meson assms associatedD divides_trans) lemma (in monoid) division_weak_partial_order [simp, intro!]: "weak_partial_order (division_rel G)" apply unfold_locales apply (simp_all add: associated_sym divides_antisym) apply (metis associated_trans) apply (metis divides_trans) by (meson associated_def divides_trans) subsubsection \Multiplication and associativity\ lemma (in monoid) mult_cong_r: assumes "b \ b'" "a \ carrier G" "b \ carrier G" "b' \ carrier G" shows "a \ b \ a \ b'" by (meson assms associated_def divides_mult_lI) lemma (in comm_monoid) mult_cong_l: assumes "a \ a'" "a \ carrier G" "a' \ carrier G" "b \ carrier G" shows "a \ b \ a' \ b" using assms m_comm mult_cong_r by auto lemma (in monoid_cancel) assoc_l_cancel: assumes "a \ carrier G" "b \ carrier G" "b' \ carrier G" "a \ b \ a \ b'" shows "b \ b'" by (meson assms associated_def divides_mult_l) lemma (in comm_monoid_cancel) assoc_r_cancel: assumes "a \ b \ a' \ b" "a \ carrier G" "a' \ carrier G" "b \ carrier G" shows "a \ a'" using assms assoc_l_cancel m_comm by presburger subsubsection \Units\ lemma (in monoid_cancel) assoc_unit_l [trans]: assumes "a \ b" and "b \ Units G" and "a \ carrier G" shows "a \ Units G" using assms by (fast elim: associatedE2) lemma (in monoid_cancel) assoc_unit_r [trans]: assumes aunit: "a \ Units G" and asc: "a \ b" and bcarr: "b \ carrier G" shows "b \ Units G" using aunit bcarr associated_sym[OF asc] by (blast intro: assoc_unit_l) lemma (in comm_monoid) Units_cong: assumes aunit: "a \ Units G" and asc: "a \ b" and bcarr: "b \ carrier G" shows "b \ Units G" using assms by (blast intro: divides_unit elim: associatedE) lemma (in monoid) Units_assoc: assumes units: "a \ Units G" "b \ Units G" shows "a \ b" using units by (fast intro: associatedI unit_divides) lemma (in monoid) Units_are_ones: "Units G {.=}\<^bsub>(division_rel G)\<^esub> {\}" proof - have "a .\\<^bsub>division_rel G\<^esub> {\}" if "a \ Units G" for a proof - have "a \ \" by (rule associatedI) (simp_all add: Units_closed that unit_divides) then show ?thesis by (simp add: elem_def) qed moreover have "\ .\\<^bsub>division_rel G\<^esub> Units G" by (simp add: equivalence.mem_imp_elem) ultimately show ?thesis by (auto simp: set_eq_def) qed lemma (in comm_monoid) Units_Lower: "Units G = Lower (division_rel G) (carrier G)" apply (auto simp add: Units_def Lower_def) apply (metis Units_one_closed unit_divides unit_factor) apply (metis Unit_eq_dividesone Units_r_inv_ex m_ac(2) one_closed) done lemma (in monoid_cancel) associated_iff: assumes "a \ carrier G" "b \ carrier G" shows "a \ b \ (\c \ Units G. a = b \ c)" using assms associatedI2' associatedD2 by auto subsubsection \Proper factors\ lemma properfactorI: fixes G (structure) assumes "a divides b" and "\(b divides a)" shows "properfactor G a b" using assms unfolding properfactor_def by simp lemma properfactorI2: fixes G (structure) assumes advdb: "a divides b" and neq: "\(a \ b)" shows "properfactor G a b" proof (rule properfactorI, rule advdb, rule notI) assume "b divides a" with advdb have "a \ b" by (rule associatedI) with neq show "False" by fast qed lemma (in comm_monoid_cancel) properfactorI3: assumes p: "p = a \ b" and nunit: "b \ Units G" and carr: "a \ carrier G" "b \ carrier G" shows "properfactor G a p" unfolding p using carr apply (intro properfactorI, fast) proof (clarsimp, elim dividesE) fix c assume ccarr: "c \ carrier G" note [simp] = carr ccarr have "a \ \ = a" by simp also assume "a = a \ b \ c" also have "\ = a \ (b \ c)" by (simp add: m_assoc) finally have "a \ \ = a \ (b \ c)" . then have rinv: "\ = b \ c" by (intro l_cancel[of "a" "\" "b \ c"], simp+) also have "\ = c \ b" by (simp add: m_comm) finally have linv: "\ = c \ b" . from ccarr linv[symmetric] rinv[symmetric] have "b \ Units G" unfolding Units_def by fastforce with nunit show False .. qed lemma properfactorE: fixes G (structure) assumes pf: "properfactor G a b" and r: "\a divides b; \(b divides a)\ \ P" shows "P" using pf unfolding properfactor_def by (fast intro: r) lemma properfactorE2: fixes G (structure) assumes pf: "properfactor G a b" and elim: "\a divides b; \(a \ b)\ \ P" shows "P" using pf unfolding properfactor_def by (fast elim: elim associatedE) lemma (in monoid) properfactor_unitE: assumes uunit: "u \ Units G" and pf: "properfactor G a u" and acarr: "a \ carrier G" shows "P" using pf unit_divides[OF uunit acarr] by (fast elim: properfactorE) lemma (in monoid) properfactor_divides: assumes pf: "properfactor G a b" shows "a divides b" using pf by (elim properfactorE) lemma (in monoid) properfactor_trans1 [trans]: assumes "a divides b" "properfactor G b c" "a \ carrier G" "c \ carrier G" shows "properfactor G a c" by (meson divides_trans properfactorE properfactorI assms) lemma (in monoid) properfactor_trans2 [trans]: assumes "properfactor G a b" "b divides c" "a \ carrier G" "b \ carrier G" shows "properfactor G a c" by (meson divides_trans properfactorE properfactorI assms) lemma properfactor_lless: fixes G (structure) shows "properfactor G = lless (division_rel G)" by (force simp: lless_def properfactor_def associated_def) lemma (in monoid) properfactor_cong_l [trans]: assumes x'x: "x' \ x" and pf: "properfactor G x y" and carr: "x \ carrier G" "x' \ carrier G" "y \ carrier G" shows "properfactor G x' y" using pf unfolding properfactor_lless proof - interpret weak_partial_order "division_rel G" .. from x'x have "x' .=\<^bsub>division_rel G\<^esub> x" by simp also assume "x \\<^bsub>division_rel G\<^esub> y" finally show "x' \\<^bsub>division_rel G\<^esub> y" by (simp add: carr) qed lemma (in monoid) properfactor_cong_r [trans]: assumes pf: "properfactor G x y" and yy': "y \ y'" and carr: "x \ carrier G" "y \ carrier G" "y' \ carrier G" shows "properfactor G x y'" using pf unfolding properfactor_lless proof - interpret weak_partial_order "division_rel G" .. assume "x \\<^bsub>division_rel G\<^esub> y" also from yy' have "y .=\<^bsub>division_rel G\<^esub> y'" by simp finally show "x \\<^bsub>division_rel G\<^esub> y'" by (simp add: carr) qed lemma (in monoid_cancel) properfactor_mult_lI [intro]: assumes ab: "properfactor G a b" and carr: "a \ carrier G" "c \ carrier G" shows "properfactor G (c \ a) (c \ b)" using ab carr by (fastforce elim: properfactorE intro: properfactorI) lemma (in monoid_cancel) properfactor_mult_l [simp]: assumes carr: "a \ carrier G" "b \ carrier G" "c \ carrier G" shows "properfactor G (c \ a) (c \ b) = properfactor G a b" using carr by (fastforce elim: properfactorE intro: properfactorI) lemma (in comm_monoid_cancel) properfactor_mult_rI [intro]: assumes ab: "properfactor G a b" and carr: "a \ carrier G" "c \ carrier G" shows "properfactor G (a \ c) (b \ c)" using ab carr by (fastforce elim: properfactorE intro: properfactorI) lemma (in comm_monoid_cancel) properfactor_mult_r [simp]: assumes carr: "a \ carrier G" "b \ carrier G" "c \ carrier G" shows "properfactor G (a \ c) (b \ c) = properfactor G a b" using carr by (fastforce elim: properfactorE intro: properfactorI) lemma (in monoid) properfactor_prod_r: assumes ab: "properfactor G a b" and carr[simp]: "a \ carrier G" "b \ carrier G" "c \ carrier G" shows "properfactor G a (b \ c)" by (intro properfactor_trans2[OF ab] divides_prod_r) simp_all lemma (in comm_monoid) properfactor_prod_l: assumes ab: "properfactor G a b" and carr[simp]: "a \ carrier G" "b \ carrier G" "c \ carrier G" shows "properfactor G a (c \ b)" by (intro properfactor_trans2[OF ab] divides_prod_l) simp_all subsection \Irreducible Elements and Primes\ subsubsection \Irreducible elements\ lemma irreducibleI: fixes G (structure) assumes "a \ Units G" and "\b. \b \ carrier G; properfactor G b a\ \ b \ Units G" shows "irreducible G a" using assms unfolding irreducible_def by blast lemma irreducibleE: fixes G (structure) assumes irr: "irreducible G a" and elim: "\a \ Units G; \b. b \ carrier G \ properfactor G b a \ b \ Units G\ \ P" shows "P" using assms unfolding irreducible_def by blast lemma irreducibleD: fixes G (structure) assumes irr: "irreducible G a" and pf: "properfactor G b a" and bcarr: "b \ carrier G" shows "b \ Units G" using assms by (fast elim: irreducibleE) lemma (in monoid_cancel) irreducible_cong [trans]: assumes "irreducible G a" "a \ a'" "a \ carrier G" "a' \ carrier G" shows "irreducible G a'" proof - have "a' divides a" by (meson \a \ a'\ associated_def) then show ?thesis by (metis (no_types) assms assoc_unit_l irreducibleE irreducibleI monoid.properfactor_trans2 monoid_axioms) qed lemma (in monoid) irreducible_prod_rI: assumes "irreducible G a" "b \ Units G" "a \ carrier G" "b \ carrier G" shows "irreducible G (a \ b)" using assms by (metis (no_types, lifting) associatedI2' irreducible_def monoid.m_closed monoid_axioms prod_unit_r properfactor_cong_r) lemma (in comm_monoid) irreducible_prod_lI: assumes birr: "irreducible G b" and aunit: "a \ Units G" and carr [simp]: "a \ carrier G" "b \ carrier G" shows "irreducible G (a \ b)" by (metis aunit birr carr irreducible_prod_rI m_comm) lemma (in comm_monoid_cancel) irreducible_prodE [elim]: assumes irr: "irreducible G (a \ b)" and carr[simp]: "a \ carrier G" "b \ carrier G" and e1: "\irreducible G a; b \ Units G\ \ P" and e2: "\a \ Units G; irreducible G b\ \ P" shows P using irr proof (elim irreducibleE) assume abnunit: "a \ b \ Units G" and isunit[rule_format]: "\ba. ba \ carrier G \ properfactor G ba (a \ b) \ ba \ Units G" show P proof (cases "a \ Units G") case aunit: True have "irreducible G b" proof (rule irreducibleI, rule notI) assume "b \ Units G" with aunit have "(a \ b) \ Units G" by fast with abnunit show "False" .. next fix c assume ccarr: "c \ carrier G" and "properfactor G c b" then have "properfactor G c (a \ b)" by (simp add: properfactor_prod_l[of c b a]) with ccarr show "c \ Units G" by (fast intro: isunit) qed with aunit show "P" by (rule e2) next case anunit: False with carr have "properfactor G b (b \ a)" by (fast intro: properfactorI3) then have bf: "properfactor G b (a \ b)" by (subst m_comm[of a b], simp+) then have bunit: "b \ Units G" by (intro isunit, simp) have "irreducible G a" proof (rule irreducibleI, rule notI) assume "a \ Units G" with bunit have "(a \ b) \ Units G" by fast with abnunit show "False" .. next fix c assume ccarr: "c \ carrier G" and "properfactor G c a" then have "properfactor G c (a \ b)" by (simp add: properfactor_prod_r[of c a b]) with ccarr show "c \ Units G" by (fast intro: isunit) qed from this bunit show "P" by (rule e1) qed qed lemma divides_irreducible_condition: assumes "irreducible G r" and "a \ carrier G" shows "a divides\<^bsub>G\<^esub> r \ a \ Units G \ a \\<^bsub>G\<^esub> r" using assms unfolding irreducible_def properfactor_def associated_def by (cases "r divides\<^bsub>G\<^esub> a", auto) subsubsection \Prime elements\ lemma primeI: fixes G (structure) assumes "p \ Units G" and "\a b. \a \ carrier G; b \ carrier G; p divides (a \ b)\ \ p divides a \ p divides b" shows "prime G p" using assms unfolding prime_def by blast lemma primeE: fixes G (structure) assumes pprime: "prime G p" and e: "\p \ Units G; \a\carrier G. \b\carrier G. p divides a \ b \ p divides a \ p divides b\ \ P" shows "P" using pprime unfolding prime_def by (blast dest: e) lemma (in comm_monoid_cancel) prime_divides: assumes carr: "a \ carrier G" "b \ carrier G" and pprime: "prime G p" and pdvd: "p divides a \ b" shows "p divides a \ p divides b" using assms by (blast elim: primeE) lemma (in monoid_cancel) prime_cong [trans]: assumes "prime G p" and pp': "p \ p'" "p \ carrier G" "p' \ carrier G" shows "prime G p'" using assms by (auto simp: prime_def assoc_unit_l) (metis pp' associated_sym divides_cong_l) lemma (in comm_monoid_cancel) prime_irreducible: \<^marker>\contributor \Paulo Emílio de Vilhena\\ assumes "prime G p" shows "irreducible G p" proof (rule irreducibleI) show "p \ Units G" using assms unfolding prime_def by simp next fix b assume A: "b \ carrier G" "properfactor G b p" then obtain c where c: "c \ carrier G" "p = b \ c" unfolding properfactor_def factor_def by auto hence "p divides c" using A assms unfolding prime_def properfactor_def by auto then obtain b' where b': "b' \ carrier G" "c = p \ b'" unfolding factor_def by auto hence "\ = b \ b'" by (metis A(1) l_cancel m_closed m_lcomm one_closed r_one c) thus "b \ Units G" using A(1) Units_one_closed b'(1) unit_factor by presburger qed lemma (in comm_monoid_cancel) prime_pow_divides_iff: assumes "p \ carrier G" "a \ carrier G" "b \ carrier G" and "prime G p" and "\ (p divides a)" shows "(p [^] (n :: nat)) divides (a \ b) \ (p [^] n) divides b" proof assume "(p [^] n) divides b" thus "(p [^] n) divides (a \ b)" using divides_prod_l[of "p [^] n" b a] assms by simp next assume "(p [^] n) divides (a \ b)" thus "(p [^] n) divides b" proof (induction n) case 0 with \b \ carrier G\ show ?case by (simp add: unit_divides) next case (Suc n) hence "(p [^] n) divides (a \ b)" and "(p [^] n) divides b" using assms(1) divides_prod_r by auto with \(p [^] (Suc n)) divides (a \ b)\ obtain c d where c: "c \ carrier G" and "b = (p [^] n) \ c" and d: "d \ carrier G" and "a \ b = (p [^] (Suc n)) \ d" using assms by blast hence "(p [^] n) \ (a \ c) = (p [^] n) \ (p \ d)" using assms by (simp add: m_assoc m_lcomm) hence "a \ c = p \ d" using c d assms(1) assms(2) l_cancel by blast with \\ (p divides a)\ and \prime G p\ have "p divides c" by (metis assms(2) c d dividesI' prime_divides) with \b = (p [^] n) \ c\ show ?case using assms(1) c by simp qed qed subsection \Factorization and Factorial Monoids\ subsubsection \Function definitions\ definition factors :: "('a, _) monoid_scheme \ 'a list \ 'a \ bool" where "factors G fs a \ (\x \ (set fs). irreducible G x) \ foldr (\\<^bsub>G\<^esub>) fs \\<^bsub>G\<^esub> = a" definition wfactors ::"('a, _) monoid_scheme \ 'a list \ 'a \ bool" where "wfactors G fs a \ (\x \ (set fs). irreducible G x) \ foldr (\\<^bsub>G\<^esub>) fs \\<^bsub>G\<^esub> \\<^bsub>G\<^esub> a" abbreviation list_assoc :: "('a, _) monoid_scheme \ 'a list \ 'a list \ bool" (infix "[\]\" 44) where "list_assoc G \ list_all2 (\\<^bsub>G\<^esub>)" definition essentially_equal :: "('a, _) monoid_scheme \ 'a list \ 'a list \ bool" where "essentially_equal G fs1 fs2 \ (\fs1'. fs1 <~~> fs1' \ fs1' [\]\<^bsub>G\<^esub> fs2)" locale factorial_monoid = comm_monoid_cancel + assumes factors_exist: "\a \ carrier G; a \ Units G\ \ \fs. set fs \ carrier G \ factors G fs a" and factors_unique: "\factors G fs a; factors G fs' a; a \ carrier G; a \ Units G; set fs \ carrier G; set fs' \ carrier G\ \ essentially_equal G fs fs'" subsubsection \Comparing lists of elements\ text \Association on lists\ lemma (in monoid) listassoc_refl [simp, intro]: assumes "set as \ carrier G" shows "as [\] as" using assms by (induct as) simp_all lemma (in monoid) listassoc_sym [sym]: assumes "as [\] bs" and "set as \ carrier G" and "set bs \ carrier G" shows "bs [\] as" using assms proof (induction as arbitrary: bs) case Cons then show ?case by (induction bs) (use associated_sym in auto) qed auto lemma (in monoid) listassoc_trans [trans]: assumes "as [\] bs" and "bs [\] cs" and "set as \ carrier G" and "set bs \ carrier G" and "set cs \ carrier G" shows "as [\] cs" using assms apply (simp add: list_all2_conv_all_nth set_conv_nth, safe) by (metis (mono_tags, lifting) associated_trans nth_mem subsetCE) lemma (in monoid_cancel) irrlist_listassoc_cong: assumes "\a\set as. irreducible G a" and "as [\] bs" and "set as \ carrier G" and "set bs \ carrier G" shows "\a\set bs. irreducible G a" using assms by (fastforce simp add: list_all2_conv_all_nth set_conv_nth intro: irreducible_cong) text \Permutations\ lemma perm_map [intro]: assumes p: "a <~~> b" shows "map f a <~~> map f b" using p by simp lemma perm_map_switch: assumes m: "map f a = map f b" and p: "b <~~> c" shows "\d. a <~~> d \ map f d = map f c" proof - from m have \length a = length b\ by (rule map_eq_imp_length_eq) from p have \mset c = mset b\ by simp then obtain p where \p permutes {.. \permute_list p b = c\ by (rule mset_eq_permutation) with \length a = length b\ have \p permutes {.. by simp moreover define d where \d = permute_list p a\ ultimately have \mset a = mset d\ \map f d = map f c\ using m \p permutes {.. \permute_list p b = c\ by (auto simp flip: permute_list_map) then show ?thesis by auto qed lemma (in monoid) perm_assoc_switch: assumes a:"as [\] bs" and p: "bs <~~> cs" shows "\bs'. as <~~> bs' \ bs' [\] cs" proof - from p have \mset cs = mset bs\ by simp then obtain p where \p permutes {.. \permute_list p bs = cs\ by (rule mset_eq_permutation) moreover define bs' where \bs' = permute_list p as\ ultimately have \as <~~> bs'\ and \bs' [\] cs\ using a by (auto simp add: list_all2_permute_list_iff list_all2_lengthD) then show ?thesis by blast qed lemma (in monoid) perm_assoc_switch_r: assumes p: "as <~~> bs" and a:"bs [\] cs" shows "\bs'. as [\] bs' \ bs' <~~> cs" using a p by (rule list_all2_reorder_left_invariance) declare perm_sym [sym] lemma perm_setP: assumes perm: "as <~~> bs" and as: "P (set as)" shows "P (set bs)" using assms by (metis set_mset_mset) lemmas (in monoid) perm_closed = perm_setP[of _ _ "\as. as \ carrier G"] lemmas (in monoid) irrlist_perm_cong = perm_setP[of _ _ "\as. \a\as. irreducible G a"] text \Essentially equal factorizations\ lemma (in monoid) essentially_equalI: assumes ex: "fs1 <~~> fs1'" "fs1' [\] fs2" shows "essentially_equal G fs1 fs2" using ex unfolding essentially_equal_def by fast lemma (in monoid) essentially_equalE: assumes ee: "essentially_equal G fs1 fs2" and e: "\fs1'. \fs1 <~~> fs1'; fs1' [\] fs2\ \ P" shows "P" using ee unfolding essentially_equal_def by (fast intro: e) lemma (in monoid) ee_refl [simp,intro]: assumes carr: "set as \ carrier G" shows "essentially_equal G as as" using carr by (fast intro: essentially_equalI) lemma (in monoid) ee_sym [sym]: assumes ee: "essentially_equal G as bs" and carr: "set as \ carrier G" "set bs \ carrier G" shows "essentially_equal G bs as" using ee proof (elim essentially_equalE) fix fs assume "as <~~> fs" "fs [\] bs" from perm_assoc_switch_r [OF this] obtain fs' where a: "as [\] fs'" and p: "fs' <~~> bs" by blast from p have "bs <~~> fs'" by (rule perm_sym) with a[symmetric] carr show ?thesis by (iprover intro: essentially_equalI perm_closed) qed lemma (in monoid) ee_trans [trans]: assumes ab: "essentially_equal G as bs" and bc: "essentially_equal G bs cs" and ascarr: "set as \ carrier G" and bscarr: "set bs \ carrier G" and cscarr: "set cs \ carrier G" shows "essentially_equal G as cs" using ab bc proof (elim essentially_equalE) fix abs bcs assume "abs [\] bs" and pb: "bs <~~> bcs" from perm_assoc_switch [OF this] obtain bs' where p: "abs <~~> bs'" and a: "bs' [\] bcs" by blast assume "as <~~> abs" with p have pp: "as <~~> bs'" by simp from pp ascarr have c1: "set bs' \ carrier G" by (rule perm_closed) from pb bscarr have c2: "set bcs \ carrier G" by (rule perm_closed) assume "bcs [\] cs" then have "bs' [\] cs" using a c1 c2 cscarr listassoc_trans by blast with pp show ?thesis by (rule essentially_equalI) qed subsubsection \Properties of lists of elements\ text \Multiplication of factors in a list\ lemma (in monoid) multlist_closed [simp, intro]: assumes ascarr: "set fs \ carrier G" shows "foldr (\) fs \ \ carrier G" using ascarr by (induct fs) simp_all lemma (in comm_monoid) multlist_dividesI: assumes "f \ set fs" and "set fs \ carrier G" shows "f divides (foldr (\) fs \)" using assms proof (induction fs) case (Cons a fs) then have f: "f \ carrier G" by blast show ?case using Cons.IH Cons.prems(1) Cons.prems(2) divides_prod_l f by auto qed auto lemma (in comm_monoid_cancel) multlist_listassoc_cong: assumes "fs [\] fs'" and "set fs \ carrier G" and "set fs' \ carrier G" shows "foldr (\) fs \ \ foldr (\) fs' \" using assms proof (induct fs arbitrary: fs') case (Cons a as fs') then show ?case proof (induction fs') case (Cons b bs) then have p: "a \ foldr (\) as \ \ b \ foldr (\) as \" by (simp add: mult_cong_l) then have "foldr (\) as \ \ foldr (\) bs \" using Cons by auto with Cons have "b \ foldr (\) as \ \ b \ foldr (\) bs \" by (simp add: mult_cong_r) then show ?case using Cons.prems(3) Cons.prems(4) monoid.associated_trans monoid_axioms p by force qed auto qed auto lemma (in comm_monoid) multlist_perm_cong: assumes prm: "as <~~> bs" and ascarr: "set as \ carrier G" shows "foldr (\) as \ = foldr (\) bs \" proof - from prm have \mset (rev as) = mset (rev bs)\ by simp moreover note one_closed ultimately have \fold (\) (rev as) \ = fold (\) (rev bs) \\ by (rule fold_permuted_eq) (use ascarr in \auto intro: m_lcomm\) then show ?thesis by (simp add: foldr_conv_fold) qed lemma (in comm_monoid_cancel) multlist_ee_cong: assumes "essentially_equal G fs fs'" and "set fs \ carrier G" and "set fs' \ carrier G" shows "foldr (\) fs \ \ foldr (\) fs' \" using assms by (metis essentially_equal_def multlist_listassoc_cong multlist_perm_cong perm_closed) subsubsection \Factorization in irreducible elements\ lemma wfactorsI: fixes G (structure) assumes "\f\set fs. irreducible G f" and "foldr (\) fs \ \ a" shows "wfactors G fs a" using assms unfolding wfactors_def by simp lemma wfactorsE: fixes G (structure) assumes wf: "wfactors G fs a" and e: "\\f\set fs. irreducible G f; foldr (\) fs \ \ a\ \ P" shows "P" using wf unfolding wfactors_def by (fast dest: e) lemma (in monoid) factorsI: assumes "\f\set fs. irreducible G f" and "foldr (\) fs \ = a" shows "factors G fs a" using assms unfolding factors_def by simp lemma factorsE: fixes G (structure) assumes f: "factors G fs a" and e: "\\f\set fs. irreducible G f; foldr (\) fs \ = a\ \ P" shows "P" using f unfolding factors_def by (simp add: e) lemma (in monoid) factors_wfactors: assumes "factors G as a" and "set as \ carrier G" shows "wfactors G as a" using assms by (blast elim: factorsE intro: wfactorsI) lemma (in monoid) wfactors_factors: assumes "wfactors G as a" and "set as \ carrier G" shows "\a'. factors G as a' \ a' \ a" using assms by (blast elim: wfactorsE intro: factorsI) lemma (in monoid) factors_closed [dest]: assumes "factors G fs a" and "set fs \ carrier G" shows "a \ carrier G" using assms by (elim factorsE, clarsimp) lemma (in monoid) nunit_factors: assumes anunit: "a \ Units G" and fs: "factors G as a" shows "length as > 0" proof - from anunit Units_one_closed have "a \ \" by auto with fs show ?thesis by (auto elim: factorsE) qed lemma (in monoid) unit_wfactors [simp]: assumes aunit: "a \ Units G" shows "wfactors G [] a" using aunit by (intro wfactorsI) (simp, simp add: Units_assoc) lemma (in comm_monoid_cancel) unit_wfactors_empty: assumes aunit: "a \ Units G" and wf: "wfactors G fs a" and carr[simp]: "set fs \ carrier G" shows "fs = []" proof (cases fs) case fs: (Cons f fs') from carr have fcarr[simp]: "f \ carrier G" and carr'[simp]: "set fs' \ carrier G" by (simp_all add: fs) from fs wf have "irreducible G f" by (simp add: wfactors_def) then have fnunit: "f \ Units G" by (fast elim: irreducibleE) from fs wf have a: "f \ foldr (\) fs' \ \ a" by (simp add: wfactors_def) note aunit also from fs wf have a: "f \ foldr (\) fs' \ \ a" by (simp add: wfactors_def) have "a \ f \ foldr (\) fs' \" by (simp add: Units_closed[OF aunit] a[symmetric]) finally have "f \ foldr (\) fs' \ \ Units G" by simp then have "f \ Units G" by (intro unit_factor[of f], simp+) with fnunit show ?thesis by contradiction qed text \Comparing wfactors\ lemma (in comm_monoid_cancel) wfactors_listassoc_cong_l: assumes fact: "wfactors G fs a" and asc: "fs [\] fs'" and carr: "a \ carrier G" "set fs \ carrier G" "set fs' \ carrier G" shows "wfactors G fs' a" proof - { from asc[symmetric] have "foldr (\) fs' \ \ foldr (\) fs \" by (simp add: multlist_listassoc_cong carr) also assume "foldr (\) fs \ \ a" finally have "foldr (\) fs' \ \ a" by (simp add: carr) } then show ?thesis using fact by (meson asc carr(2) carr(3) irrlist_listassoc_cong wfactors_def) qed lemma (in comm_monoid) wfactors_perm_cong_l: assumes "wfactors G fs a" and "fs <~~> fs'" and "set fs \ carrier G" shows "wfactors G fs' a" using assms irrlist_perm_cong multlist_perm_cong wfactors_def by fastforce lemma (in comm_monoid_cancel) wfactors_ee_cong_l [trans]: assumes ee: "essentially_equal G as bs" and bfs: "wfactors G bs b" and carr: "b \ carrier G" "set as \ carrier G" "set bs \ carrier G" shows "wfactors G as b" using ee proof (elim essentially_equalE) fix fs assume prm: "as <~~> fs" with carr have fscarr: "set fs \ carrier G" using perm_closed by blast note bfs also assume [symmetric]: "fs [\] bs" also (wfactors_listassoc_cong_l) have \mset fs = mset as\ using prm by simp finally (wfactors_perm_cong_l) show "wfactors G as b" by (simp add: carr fscarr) qed lemma (in monoid) wfactors_cong_r [trans]: assumes fac: "wfactors G fs a" and aa': "a \ a'" and carr[simp]: "a \ carrier G" "a' \ carrier G" "set fs \ carrier G" shows "wfactors G fs a'" using fac proof (elim wfactorsE, intro wfactorsI) assume "foldr (\) fs \ \ a" also note aa' finally show "foldr (\) fs \ \ a'" by simp qed subsubsection \Essentially equal factorizations\ lemma (in comm_monoid_cancel) unitfactor_ee: assumes uunit: "u \ Units G" and carr: "set as \ carrier G" shows "essentially_equal G (as[0 := (as!0 \ u)]) as" (is "essentially_equal G ?as' as") proof - have "as[0 := as ! 0 \ u] [\] as" proof (cases as) case (Cons a as') then show ?thesis using associatedI2 carr uunit by auto qed auto then show ?thesis using essentially_equal_def by blast qed lemma (in comm_monoid_cancel) factors_cong_unit: assumes u: "u \ Units G" and a: "a \ Units G" and afs: "factors G as a" and ascarr: "set as \ carrier G" shows "factors G (as[0 := (as!0 \ u)]) (a \ u)" (is "factors G ?as' ?a'") proof (cases as) case Nil then show ?thesis using afs a nunit_factors by auto next case (Cons b bs) have *: "\f\set as. irreducible G f" "foldr (\) as \ = a" using afs by (auto simp: factors_def) show ?thesis proof (intro factorsI) show "foldr (\) (as[0 := as ! 0 \ u]) \ = a \ u" using Cons u ascarr * by (auto simp add: m_ac Units_closed) show "\f\set (as[0 := as ! 0 \ u]). irreducible G f" using Cons u ascarr * by (force intro: irreducible_prod_rI) qed qed lemma (in comm_monoid) perm_wfactorsD: assumes prm: "as <~~> bs" and afs: "wfactors G as a" and bfs: "wfactors G bs b" and [simp]: "a \ carrier G" "b \ carrier G" and ascarr [simp]: "set as \ carrier G" shows "a \ b" using afs bfs proof (elim wfactorsE) from prm have [simp]: "set bs \ carrier G" by (simp add: perm_closed) assume "foldr (\) as \ \ a" then have "a \ foldr (\) as \" by (simp add: associated_sym) also from prm have "foldr (\) as \ = foldr (\) bs \" by (rule multlist_perm_cong, simp) also assume "foldr (\) bs \ \ b" finally show "a \ b" by simp qed lemma (in comm_monoid_cancel) listassoc_wfactorsD: assumes assoc: "as [\] bs" and afs: "wfactors G as a" and bfs: "wfactors G bs b" and [simp]: "a \ carrier G" "b \ carrier G" and [simp]: "set as \ carrier G" "set bs \ carrier G" shows "a \ b" using afs bfs proof (elim wfactorsE) assume "foldr (\) as \ \ a" then have "a \ foldr (\) as \" by (simp add: associated_sym) also from assoc have "foldr (\) as \ \ foldr (\) bs \" by (rule multlist_listassoc_cong, simp+) also assume "foldr (\) bs \ \ b" finally show "a \ b" by simp qed lemma (in comm_monoid_cancel) ee_wfactorsD: assumes ee: "essentially_equal G as bs" and afs: "wfactors G as a" and bfs: "wfactors G bs b" and [simp]: "a \ carrier G" "b \ carrier G" and ascarr[simp]: "set as \ carrier G" and bscarr[simp]: "set bs \ carrier G" shows "a \ b" using ee proof (elim essentially_equalE) fix fs assume prm: "as <~~> fs" then have as'carr[simp]: "set fs \ carrier G" by (simp add: perm_closed) from afs prm have afs': "wfactors G fs a" by (rule wfactors_perm_cong_l) simp assume "fs [\] bs" from this afs' bfs show "a \ b" by (rule listassoc_wfactorsD) simp_all qed lemma (in comm_monoid_cancel) ee_factorsD: assumes ee: "essentially_equal G as bs" and afs: "factors G as a" and bfs:"factors G bs b" and "set as \ carrier G" "set bs \ carrier G" shows "a \ b" using assms by (blast intro: factors_wfactors dest: ee_wfactorsD) lemma (in factorial_monoid) ee_factorsI: assumes ab: "a \ b" and afs: "factors G as a" and anunit: "a \ Units G" and bfs: "factors G bs b" and bnunit: "b \ Units G" and ascarr: "set as \ carrier G" and bscarr: "set bs \ carrier G" shows "essentially_equal G as bs" proof - note carr[simp] = factors_closed[OF afs ascarr] ascarr[THEN subsetD] factors_closed[OF bfs bscarr] bscarr[THEN subsetD] from ab carr obtain u where uunit: "u \ Units G" and a: "a = b \ u" by (elim associatedE2) from uunit bscarr have ee: "essentially_equal G (bs[0 := (bs!0 \ u)]) bs" (is "essentially_equal G ?bs' bs") by (rule unitfactor_ee) from bscarr uunit have bs'carr: "set ?bs' \ carrier G" by (cases bs) (simp_all add: Units_closed) from uunit bnunit bfs bscarr have fac: "factors G ?bs' (b \ u)" by (rule factors_cong_unit) from afs fac[simplified a[symmetric]] ascarr bs'carr anunit have "essentially_equal G as ?bs'" by (blast intro: factors_unique) also note ee finally show "essentially_equal G as bs" by (simp add: ascarr bscarr bs'carr) qed lemma (in factorial_monoid) ee_wfactorsI: assumes asc: "a \ b" and asf: "wfactors G as a" and bsf: "wfactors G bs b" and acarr[simp]: "a \ carrier G" and bcarr[simp]: "b \ carrier G" and ascarr[simp]: "set as \ carrier G" and bscarr[simp]: "set bs \ carrier G" shows "essentially_equal G as bs" using assms proof (cases "a \ Units G") case aunit: True also note asc finally have bunit: "b \ Units G" by simp from aunit asf ascarr have e: "as = []" by (rule unit_wfactors_empty) from bunit bsf bscarr have e': "bs = []" by (rule unit_wfactors_empty) have "essentially_equal G [] []" by (fast intro: essentially_equalI) then show ?thesis by (simp add: e e') next case anunit: False have bnunit: "b \ Units G" proof clarify assume "b \ Units G" also note asc[symmetric] finally have "a \ Units G" by simp with anunit show False .. qed from wfactors_factors[OF asf ascarr] obtain a' where fa': "factors G as a'" and a': "a' \ a" by blast from fa' ascarr have a'carr[simp]: "a' \ carrier G" by fast have a'nunit: "a' \ Units G" proof clarify assume "a' \ Units G" also note a' finally have "a \ Units G" by simp with anunit show "False" .. qed from wfactors_factors[OF bsf bscarr] obtain b' where fb': "factors G bs b'" and b': "b' \ b" by blast from fb' bscarr have b'carr[simp]: "b' \ carrier G" by fast have b'nunit: "b' \ Units G" proof clarify assume "b' \ Units G" also note b' finally have "b \ Units G" by simp with bnunit show False .. qed note a' also note asc also note b'[symmetric] finally have "a' \ b'" by simp from this fa' a'nunit fb' b'nunit ascarr bscarr show "essentially_equal G as bs" by (rule ee_factorsI) qed lemma (in factorial_monoid) ee_wfactors: assumes asf: "wfactors G as a" and bsf: "wfactors G bs b" and acarr: "a \ carrier G" and bcarr: "b \ carrier G" and ascarr: "set as \ carrier G" and bscarr: "set bs \ carrier G" shows asc: "a \ b = essentially_equal G as bs" using assms by (fast intro: ee_wfactorsI ee_wfactorsD) lemma (in factorial_monoid) wfactors_exist [intro, simp]: assumes acarr[simp]: "a \ carrier G" shows "\fs. set fs \ carrier G \ wfactors G fs a" proof (cases "a \ Units G") case True then have "wfactors G [] a" by (rule unit_wfactors) then show ?thesis by (intro exI) force next case False with factors_exist [OF acarr] obtain fs where fscarr: "set fs \ carrier G" and f: "factors G fs a" by blast from f have "wfactors G fs a" by (rule factors_wfactors) fact with fscarr show ?thesis by fast qed lemma (in monoid) wfactors_prod_exists [intro, simp]: assumes "\a \ set as. irreducible G a" and "set as \ carrier G" shows "\a. a \ carrier G \ wfactors G as a" unfolding wfactors_def using assms by blast lemma (in factorial_monoid) wfactors_unique: assumes "wfactors G fs a" and "wfactors G fs' a" and "a \ carrier G" and "set fs \ carrier G" and "set fs' \ carrier G" shows "essentially_equal G fs fs'" using assms by (fast intro: ee_wfactorsI[of a a]) lemma (in monoid) factors_mult_single: assumes "irreducible G a" and "factors G fb b" and "a \ carrier G" shows "factors G (a # fb) (a \ b)" using assms unfolding factors_def by simp lemma (in monoid_cancel) wfactors_mult_single: assumes f: "irreducible G a" "wfactors G fb b" "a \ carrier G" "b \ carrier G" "set fb \ carrier G" shows "wfactors G (a # fb) (a \ b)" using assms unfolding wfactors_def by (simp add: mult_cong_r) lemma (in monoid) factors_mult: assumes factors: "factors G fa a" "factors G fb b" and ascarr: "set fa \ carrier G" and bscarr: "set fb \ carrier G" shows "factors G (fa @ fb) (a \ b)" proof - have "foldr (\) (fa @ fb) \ = foldr (\) fa \ \ foldr (\) fb \" if "set fa \ carrier G" "Ball (set fa) (irreducible G)" using that bscarr by (induct fa) (simp_all add: m_assoc) then show ?thesis using assms unfolding factors_def by force qed lemma (in comm_monoid_cancel) wfactors_mult [intro]: assumes asf: "wfactors G as a" and bsf:"wfactors G bs b" and acarr: "a \ carrier G" and bcarr: "b \ carrier G" and ascarr: "set as \ carrier G" and bscarr:"set bs \ carrier G" shows "wfactors G (as @ bs) (a \ b)" using wfactors_factors[OF asf ascarr] and wfactors_factors[OF bsf bscarr] proof clarsimp fix a' b' assume asf': "factors G as a'" and a'a: "a' \ a" and bsf': "factors G bs b'" and b'b: "b' \ b" from asf' have a'carr: "a' \ carrier G" by (rule factors_closed) fact from bsf' have b'carr: "b' \ carrier G" by (rule factors_closed) fact note carr = acarr bcarr a'carr b'carr ascarr bscarr from asf' bsf' have "factors G (as @ bs) (a' \ b')" by (rule factors_mult) fact+ with carr have abf': "wfactors G (as @ bs) (a' \ b')" by (intro factors_wfactors) simp_all also from b'b carr have trb: "a' \ b' \ a' \ b" by (intro mult_cong_r) also from a'a carr have tra: "a' \ b \ a \ b" by (intro mult_cong_l) finally show "wfactors G (as @ bs) (a \ b)" by (simp add: carr) qed lemma (in comm_monoid) factors_dividesI: assumes "factors G fs a" and "f \ set fs" and "set fs \ carrier G" shows "f divides a" using assms by (fast elim: factorsE intro: multlist_dividesI) lemma (in comm_monoid) wfactors_dividesI: assumes p: "wfactors G fs a" and fscarr: "set fs \ carrier G" and acarr: "a \ carrier G" and f: "f \ set fs" shows "f divides a" using wfactors_factors[OF p fscarr] proof clarsimp fix a' assume fsa': "factors G fs a'" and a'a: "a' \ a" with fscarr have a'carr: "a' \ carrier G" by (simp add: factors_closed) from fsa' fscarr f have "f divides a'" by (fast intro: factors_dividesI) also note a'a finally show "f divides a" by (simp add: f fscarr[THEN subsetD] acarr a'carr) qed subsubsection \Factorial monoids and wfactors\ lemma (in comm_monoid_cancel) factorial_monoidI: assumes wfactors_exists: "\a. \ a \ carrier G; a \ Units G \ \ \fs. set fs \ carrier G \ wfactors G fs a" and wfactors_unique: "\a fs fs'. \a \ carrier G; set fs \ carrier G; set fs' \ carrier G; wfactors G fs a; wfactors G fs' a\ \ essentially_equal G fs fs'" shows "factorial_monoid G" proof fix a assume acarr: "a \ carrier G" and anunit: "a \ Units G" from wfactors_exists[OF acarr anunit] obtain as where ascarr: "set as \ carrier G" and afs: "wfactors G as a" by blast from wfactors_factors [OF afs ascarr] obtain a' where afs': "factors G as a'" and a'a: "a' \ a" by blast from afs' ascarr have a'carr: "a' \ carrier G" by fast have a'nunit: "a' \ Units G" proof clarify assume "a' \ Units G" also note a'a finally have "a \ Units G" by (simp add: acarr) with anunit show False .. qed from a'carr acarr a'a obtain u where uunit: "u \ Units G" and a': "a' = a \ u" by (blast elim: associatedE2) note [simp] = acarr Units_closed[OF uunit] Units_inv_closed[OF uunit] have "a = a \ \" by simp also have "\ = a \ (u \ inv u)" by (simp add: uunit) also have "\ = a' \ inv u" by (simp add: m_assoc[symmetric] a'[symmetric]) finally have a: "a = a' \ inv u" . from ascarr uunit have cr: "set (as[0:=(as!0 \ inv u)]) \ carrier G" by (cases as) auto from afs' uunit a'nunit acarr ascarr have "factors G (as[0:=(as!0 \ inv u)]) a" by (simp add: a factors_cong_unit) with cr show "\fs. set fs \ carrier G \ factors G fs a" by fast qed (blast intro: factors_wfactors wfactors_unique) subsection \Factorizations as Multisets\ text \Gives useful operations like intersection\ (* FIXME: use class_of x instead of closure_of {x} *) abbreviation "assocs G x \ eq_closure_of (division_rel G) {x}" definition "fmset G as = mset (map (assocs G) as)" text \Helper lemmas\ lemma (in monoid) assocs_repr_independence: assumes "y \ assocs G x" "x \ carrier G" shows "assocs G x = assocs G y" using assms by (simp add: eq_closure_of_def elem_def) (use associated_sym associated_trans in \blast+\) lemma (in monoid) assocs_self: assumes "x \ carrier G" shows "x \ assocs G x" using assms by (fastforce intro: closure_ofI2) lemma (in monoid) assocs_repr_independenceD: assumes repr: "assocs G x = assocs G y" and ycarr: "y \ carrier G" shows "y \ assocs G x" unfolding repr using ycarr by (intro assocs_self) lemma (in comm_monoid) assocs_assoc: assumes "a \ assocs G b" "b \ carrier G" shows "a \ b" using assms by (elim closure_ofE2) simp lemmas (in comm_monoid) assocs_eqD = assocs_repr_independenceD[THEN assocs_assoc] subsubsection \Comparing multisets\ lemma (in monoid) fmset_perm_cong: assumes prm: "as <~~> bs" shows "fmset G as = fmset G bs" using perm_map[OF prm] unfolding fmset_def by blast lemma (in comm_monoid_cancel) eqc_listassoc_cong: assumes "as [\] bs" and "set as \ carrier G" and "set bs \ carrier G" shows "map (assocs G) as = map (assocs G) bs" using assms proof (induction as arbitrary: bs) case Nil then show ?case by simp next case (Cons a as) then show ?case proof (clarsimp simp add: Cons_eq_map_conv list_all2_Cons1) fix z zs assume zzs: "a \ carrier G" "set as \ carrier G" "bs = z # zs" "a \ z" "as [\] zs" "z \ carrier G" "set zs \ carrier G" then show "assocs G a = assocs G z" apply (simp add: eq_closure_of_def elem_def) using \a \ carrier G\ \z \ carrier G\ \a \ z\ associated_sym associated_trans by blast+ qed qed lemma (in comm_monoid_cancel) fmset_listassoc_cong: assumes "as [\] bs" and "set as \ carrier G" and "set bs \ carrier G" shows "fmset G as = fmset G bs" using assms unfolding fmset_def by (simp add: eqc_listassoc_cong) lemma (in comm_monoid_cancel) ee_fmset: assumes ee: "essentially_equal G as bs" and ascarr: "set as \ carrier G" and bscarr: "set bs \ carrier G" shows "fmset G as = fmset G bs" using ee thm essentially_equal_def proof (elim essentially_equalE) fix as' assume prm: "as <~~> as'" from prm ascarr have as'carr: "set as' \ carrier G" by (rule perm_closed) from prm have "fmset G as = fmset G as'" by (rule fmset_perm_cong) also assume "as' [\] bs" with as'carr bscarr have "fmset G as' = fmset G bs" by (simp add: fmset_listassoc_cong) finally show "fmset G as = fmset G bs" . qed lemma (in comm_monoid_cancel) fmset_ee: assumes mset: "fmset G as = fmset G bs" and ascarr: "set as \ carrier G" and bscarr: "set bs \ carrier G" shows "essentially_equal G as bs" proof - from mset have "mset (map (assocs G) bs) = mset (map (assocs G) as)" by (simp add: fmset_def) then obtain p where \p permutes {.. \permute_list p (map (assocs G) as) = map (assocs G) bs\ by (rule mset_eq_permutation) then have \p permutes {.. \map (assocs G) (permute_list p as) = map (assocs G) bs\ by (simp_all add: permute_list_map) moreover define as' where \as' = permute_list p as\ ultimately have tp: "as <~~> as'" and tm: "map (assocs G) as' = map (assocs G) bs" by simp_all from tp show ?thesis proof (rule essentially_equalI) from tm tp ascarr have as'carr: "set as' \ carrier G" using perm_closed by blast from tm as'carr[THEN subsetD] bscarr[THEN subsetD] show "as' [\] bs" by (induct as' arbitrary: bs) (simp, fastforce dest: assocs_eqD[THEN associated_sym]) qed qed lemma (in comm_monoid_cancel) ee_is_fmset: assumes "set as \ carrier G" and "set bs \ carrier G" shows "essentially_equal G as bs = (fmset G as = fmset G bs)" using assms by (fast intro: ee_fmset fmset_ee) subsubsection \Interpreting multisets as factorizations\ lemma (in monoid) mset_fmsetEx: assumes elems: "\X. X \ set_mset Cs \ \x. P x \ X = assocs G x" shows "\cs. (\c \ set cs. P c) \ fmset G cs = Cs" proof - from surjE[OF surj_mset] obtain Cs' where Cs: "Cs = mset Cs'" by blast have "\cs. (\c \ set cs. P c) \ mset (map (assocs G) cs) = Cs" using elems unfolding Cs proof (induction Cs') case (Cons a Cs') then obtain c cs where csP: "\x\set cs. P x" and mset: "mset (map (assocs G) cs) = mset Cs'" and cP: "P c" and a: "a = assocs G c" by force then have tP: "\x\set (c#cs). P x" by simp show ?case using tP mset a by fastforce qed auto then show ?thesis by (simp add: fmset_def) qed lemma (in monoid) mset_wfactorsEx: assumes elems: "\X. X \ set_mset Cs \ \x. (x \ carrier G \ irreducible G x) \ X = assocs G x" shows "\c cs. c \ carrier G \ set cs \ carrier G \ wfactors G cs c \ fmset G cs = Cs" proof - have "\cs. (\c\set cs. c \ carrier G \ irreducible G c) \ fmset G cs = Cs" by (intro mset_fmsetEx, rule elems) then obtain cs where p[rule_format]: "\c\set cs. c \ carrier G \ irreducible G c" and Cs[symmetric]: "fmset G cs = Cs" by auto from p have cscarr: "set cs \ carrier G" by fast from p have "\c. c \ carrier G \ wfactors G cs c" by (intro wfactors_prod_exists) auto then obtain c where ccarr: "c \ carrier G" and cfs: "wfactors G cs c" by auto with cscarr Cs show ?thesis by fast qed subsubsection \Multiplication on multisets\ lemma (in factorial_monoid) mult_wfactors_fmset: assumes afs: "wfactors G as a" and bfs: "wfactors G bs b" and cfs: "wfactors G cs (a \ b)" and carr: "a \ carrier G" "b \ carrier G" "set as \ carrier G" "set bs \ carrier G" "set cs \ carrier G" shows "fmset G cs = fmset G as + fmset G bs" proof - from assms have "wfactors G (as @ bs) (a \ b)" by (intro wfactors_mult) with carr cfs have "essentially_equal G cs (as@bs)" by (intro ee_wfactorsI[of "a\b" "a\b"]) simp_all with carr have "fmset G cs = fmset G (as@bs)" by (intro ee_fmset) simp_all also have "fmset G (as@bs) = fmset G as + fmset G bs" by (simp add: fmset_def) finally show "fmset G cs = fmset G as + fmset G bs" . qed lemma (in factorial_monoid) mult_factors_fmset: assumes afs: "factors G as a" and bfs: "factors G bs b" and cfs: "factors G cs (a \ b)" and "set as \ carrier G" "set bs \ carrier G" "set cs \ carrier G" shows "fmset G cs = fmset G as + fmset G bs" using assms by (blast intro: factors_wfactors mult_wfactors_fmset) lemma (in comm_monoid_cancel) fmset_wfactors_mult: assumes mset: "fmset G cs = fmset G as + fmset G bs" and carr: "a \ carrier G" "b \ carrier G" "c \ carrier G" "set as \ carrier G" "set bs \ carrier G" "set cs \ carrier G" and fs: "wfactors G as a" "wfactors G bs b" "wfactors G cs c" shows "c \ a \ b" proof - from carr fs have m: "wfactors G (as @ bs) (a \ b)" by (intro wfactors_mult) from mset have "fmset G cs = fmset G (as@bs)" by (simp add: fmset_def) then have "essentially_equal G cs (as@bs)" by (rule fmset_ee) (simp_all add: carr) then show "c \ a \ b" by (rule ee_wfactorsD[of "cs" "as@bs"]) (simp_all add: assms m) qed subsubsection \Divisibility on multisets\ lemma (in factorial_monoid) divides_fmsubset: assumes ab: "a divides b" and afs: "wfactors G as a" and bfs: "wfactors G bs b" and carr: "a \ carrier G" "b \ carrier G" "set as \ carrier G" "set bs \ carrier G" shows "fmset G as \# fmset G bs" using ab proof (elim dividesE) fix c assume ccarr: "c \ carrier G" from wfactors_exist [OF this] obtain cs where cscarr: "set cs \ carrier G" and cfs: "wfactors G cs c" by blast note carr = carr ccarr cscarr assume "b = a \ c" with afs bfs cfs carr have "fmset G bs = fmset G as + fmset G cs" by (intro mult_wfactors_fmset[OF afs cfs]) simp_all then show ?thesis by simp qed lemma (in comm_monoid_cancel) fmsubset_divides: assumes msubset: "fmset G as \# fmset G bs" and afs: "wfactors G as a" and bfs: "wfactors G bs b" and acarr: "a \ carrier G" and bcarr: "b \ carrier G" and ascarr: "set as \ carrier G" and bscarr: "set bs \ carrier G" shows "a divides b" proof - from afs have airr: "\a \ set as. irreducible G a" by (fast elim: wfactorsE) from bfs have birr: "\b \ set bs. irreducible G b" by (fast elim: wfactorsE) have "\c cs. c \ carrier G \ set cs \ carrier G \ wfactors G cs c \ fmset G cs = fmset G bs - fmset G as" proof (intro mset_wfactorsEx, simp) fix X assume "X \# fmset G bs - fmset G as" then have "X \# fmset G bs" by (rule in_diffD) then have "X \ set (map (assocs G) bs)" by (simp add: fmset_def) then have "\x. x \ set bs \ X = assocs G x" by (induct bs) auto then obtain x where xbs: "x \ set bs" and X: "X = assocs G x" by auto with bscarr have xcarr: "x \ carrier G" by fast from xbs birr have xirr: "irreducible G x" by simp from xcarr and xirr and X show "\x. x \ carrier G \ irreducible G x \ X = assocs G x" by fast qed then obtain c cs where ccarr: "c \ carrier G" and cscarr: "set cs \ carrier G" and csf: "wfactors G cs c" and csmset: "fmset G cs = fmset G bs - fmset G as" by auto from csmset msubset have "fmset G bs = fmset G as + fmset G cs" by (simp add: multiset_eq_iff subseteq_mset_def) then have basc: "b \ a \ c" by (rule fmset_wfactors_mult) fact+ then show ?thesis proof (elim associatedE2) fix u assume "u \ Units G" "b = a \ c \ u" with acarr ccarr show "a divides b" by (fast intro: dividesI[of "c \ u"] m_assoc) qed (simp_all add: acarr bcarr ccarr) qed lemma (in factorial_monoid) divides_as_fmsubset: assumes "wfactors G as a" and "wfactors G bs b" and "a \ carrier G" and "b \ carrier G" and "set as \ carrier G" and "set bs \ carrier G" shows "a divides b = (fmset G as \# fmset G bs)" using assms by (blast intro: divides_fmsubset fmsubset_divides) text \Proper factors on multisets\ lemma (in factorial_monoid) fmset_properfactor: assumes asubb: "fmset G as \# fmset G bs" and anb: "fmset G as \ fmset G bs" and "wfactors G as a" and "wfactors G bs b" and "a \ carrier G" and "b \ carrier G" and "set as \ carrier G" and "set bs \ carrier G" shows "properfactor G a b" proof (rule properfactorI) show "a divides b" using assms asubb fmsubset_divides by blast show "\ b divides a" by (meson anb assms asubb factorial_monoid.divides_fmsubset factorial_monoid_axioms subset_mset.antisym) qed lemma (in factorial_monoid) properfactor_fmset: assumes "properfactor G a b" and "wfactors G as a" and "wfactors G bs b" and "a \ carrier G" and "b \ carrier G" and "set as \ carrier G" and "set bs \ carrier G" shows "fmset G as \# fmset G bs" using assms by (meson divides_as_fmsubset properfactor_divides) lemma (in factorial_monoid) properfactor_fmset_ne: assumes pf: "properfactor G a b" and "wfactors G as a" and "wfactors G bs b" and "a \ carrier G" and "b \ carrier G" and "set as \ carrier G" and "set bs \ carrier G" shows "fmset G as \ fmset G bs" using properfactorE [OF pf] assms divides_as_fmsubset by force subsection \Irreducible Elements are Prime\ lemma (in factorial_monoid) irreducible_prime: assumes pirr: "irreducible G p" and pcarr: "p \ carrier G" shows "prime G p" using pirr proof (elim irreducibleE, intro primeI) fix a b assume acarr: "a \ carrier G" and bcarr: "b \ carrier G" and pdvdab: "p divides (a \ b)" and pnunit: "p \ Units G" assume irreduc[rule_format]: "\b. b \ carrier G \ properfactor G b p \ b \ Units G" from pdvdab obtain c where ccarr: "c \ carrier G" and abpc: "a \ b = p \ c" by (rule dividesE) obtain as where ascarr: "set as \ carrier G" and afs: "wfactors G as a" using wfactors_exist [OF acarr] by blast obtain bs where bscarr: "set bs \ carrier G" and bfs: "wfactors G bs b" using wfactors_exist [OF bcarr] by blast obtain cs where cscarr: "set cs \ carrier G" and cfs: "wfactors G cs c" using wfactors_exist [OF ccarr] by blast note carr[simp] = pcarr acarr bcarr ccarr ascarr bscarr cscarr from pirr cfs abpc have "wfactors G (p # cs) (a \ b)" by (simp add: wfactors_mult_single) moreover have "wfactors G (as @ bs) (a \ b)" by (rule wfactors_mult [OF afs bfs]) fact+ ultimately have "essentially_equal G (p # cs) (as @ bs)" by (rule wfactors_unique) simp+ then obtain ds where "p # cs <~~> ds" and dsassoc: "ds [\] (as @ bs)" by (fast elim: essentially_equalE) then have "p \ set ds" by (metis \mset (p # cs) = mset ds\ insert_iff list.set(2) perm_set_eq) with dsassoc obtain p' where "p' \ set (as@bs)" and pp': "p \ p'" unfolding list_all2_conv_all_nth set_conv_nth by force then consider "p' \ set as" | "p' \ set bs" by auto then show "p divides a \ p divides b" using afs bfs divides_cong_l pp' wfactors_dividesI by (meson acarr ascarr bcarr bscarr pcarr) qed \ \A version using \<^const>\factors\, more complicated\ lemma (in factorial_monoid) factors_irreducible_prime: assumes pirr: "irreducible G p" and pcarr: "p \ carrier G" shows "prime G p" proof (rule primeI) show "p \ Units G" by (meson irreducibleE pirr) have irreduc: "\b. \b \ carrier G; properfactor G b p\ \ b \ Units G" using pirr by (auto simp: irreducible_def) show "p divides a \ p divides b" if acarr: "a \ carrier G" and bcarr: "b \ carrier G" and pdvdab: "p divides (a \ b)" for a b proof - from pdvdab obtain c where ccarr: "c \ carrier G" and abpc: "a \ b = p \ c" by (rule dividesE) note [simp] = pcarr acarr bcarr ccarr show "p divides a \ p divides b" proof (cases "a \ Units G") case True then have "p divides b" by (metis acarr associatedI2' associated_def bcarr divides_trans m_comm pcarr pdvdab) then show ?thesis .. next case anunit: False show ?thesis proof (cases "b \ Units G") case True then have "p divides a" by (meson acarr bcarr divides_unit irreducible_prime pcarr pdvdab pirr prime_def) then show ?thesis .. next case bnunit: False then have cnunit: "c \ Units G" by (metis abpc acarr anunit bcarr ccarr irreducible_prodE irreducible_prod_rI pcarr pirr) then have abnunit: "a \ b \ Units G" using acarr anunit bcarr unit_factor by blast obtain as where ascarr: "set as \ carrier G" and afac: "factors G as a" using factors_exist [OF acarr anunit] by blast obtain bs where bscarr: "set bs \ carrier G" and bfac: "factors G bs b" using factors_exist [OF bcarr bnunit] by blast obtain cs where cscarr: "set cs \ carrier G" and cfac: "factors G cs c" using factors_exist [OF ccarr cnunit] by auto note [simp] = ascarr bscarr cscarr from pirr cfac abpc have abfac': "factors G (p # cs) (a \ b)" by (simp add: factors_mult_single) from afac and bfac have "factors G (as @ bs) (a \ b)" by (rule factors_mult) fact+ with abfac' have "essentially_equal G (p # cs) (as @ bs)" using abnunit factors_unique by auto then obtain ds where "p # cs <~~> ds" and dsassoc: "ds [\] (as @ bs)" by (fast elim: essentially_equalE) then have "p \ set ds" by (metis list.set_intros(1) set_mset_mset) with dsassoc obtain p' where "p' \ set (as@bs)" and pp': "p \ p'" unfolding list_all2_conv_all_nth set_conv_nth by force then consider "p' \ set as" | "p' \ set bs" by auto then show "p divides a \ p divides b" by (meson afac bfac divides_cong_l factors_dividesI pp' ascarr bscarr pcarr) qed qed qed qed subsection \Greatest Common Divisors and Lowest Common Multiples\ subsubsection \Definitions\ definition isgcd :: "[('a,_) monoid_scheme, 'a, 'a, 'a] \ bool" ("(_ gcdof\ _ _)" [81,81,81] 80) where "x gcdof\<^bsub>G\<^esub> a b \ x divides\<^bsub>G\<^esub> a \ x divides\<^bsub>G\<^esub> b \ (\y\carrier G. (y divides\<^bsub>G\<^esub> a \ y divides\<^bsub>G\<^esub> b \ y divides\<^bsub>G\<^esub> x))" definition islcm :: "[_, 'a, 'a, 'a] \ bool" ("(_ lcmof\ _ _)" [81,81,81] 80) where "x lcmof\<^bsub>G\<^esub> a b \ a divides\<^bsub>G\<^esub> x \ b divides\<^bsub>G\<^esub> x \ (\y\carrier G. (a divides\<^bsub>G\<^esub> y \ b divides\<^bsub>G\<^esub> y \ x divides\<^bsub>G\<^esub> y))" definition somegcd :: "('a,_) monoid_scheme \ 'a \ 'a \ 'a" where "somegcd G a b = (SOME x. x \ carrier G \ x gcdof\<^bsub>G\<^esub> a b)" definition somelcm :: "('a,_) monoid_scheme \ 'a \ 'a \ 'a" where "somelcm G a b = (SOME x. x \ carrier G \ x lcmof\<^bsub>G\<^esub> a b)" definition "SomeGcd G A = Lattice.inf (division_rel G) A" locale gcd_condition_monoid = comm_monoid_cancel + assumes gcdof_exists: "\a \ carrier G; b \ carrier G\ \ \c. c \ carrier G \ c gcdof a b" locale primeness_condition_monoid = comm_monoid_cancel + assumes irreducible_prime: "\a \ carrier G; irreducible G a\ \ prime G a" locale divisor_chain_condition_monoid = comm_monoid_cancel + assumes division_wellfounded: "wf {(x, y). x \ carrier G \ y \ carrier G \ properfactor G x y}" subsubsection \Connections to \texttt{Lattice.thy}\ lemma gcdof_greatestLower: fixes G (structure) assumes carr[simp]: "a \ carrier G" "b \ carrier G" shows "(x \ carrier G \ x gcdof a b) = greatest (division_rel G) x (Lower (division_rel G) {a, b})" by (auto simp: isgcd_def greatest_def Lower_def elem_def) lemma lcmof_leastUpper: fixes G (structure) assumes carr[simp]: "a \ carrier G" "b \ carrier G" shows "(x \ carrier G \ x lcmof a b) = least (division_rel G) x (Upper (division_rel G) {a, b})" by (auto simp: islcm_def least_def Upper_def elem_def) lemma somegcd_meet: fixes G (structure) assumes carr: "a \ carrier G" "b \ carrier G" shows "somegcd G a b = meet (division_rel G) a b" by (simp add: somegcd_def meet_def inf_def gcdof_greatestLower[OF carr]) lemma (in monoid) isgcd_divides_l: assumes "a divides b" and "a \ carrier G" "b \ carrier G" shows "a gcdof a b" using assms unfolding isgcd_def by fast lemma (in monoid) isgcd_divides_r: assumes "b divides a" and "a \ carrier G" "b \ carrier G" shows "b gcdof a b" using assms unfolding isgcd_def by fast subsubsection \Existence of gcd and lcm\ lemma (in factorial_monoid) gcdof_exists: assumes acarr: "a \ carrier G" and bcarr: "b \ carrier G" shows "\c. c \ carrier G \ c gcdof a b" proof - from wfactors_exist [OF acarr] obtain as where ascarr: "set as \ carrier G" and afs: "wfactors G as a" by blast from afs have airr: "\a \ set as. irreducible G a" by (fast elim: wfactorsE) from wfactors_exist [OF bcarr] obtain bs where bscarr: "set bs \ carrier G" and bfs: "wfactors G bs b" by blast from bfs have birr: "\b \ set bs. irreducible G b" by (fast elim: wfactorsE) have "\c cs. c \ carrier G \ set cs \ carrier G \ wfactors G cs c \ fmset G cs = fmset G as \# fmset G bs" proof (intro mset_wfactorsEx) fix X assume "X \# fmset G as \# fmset G bs" then have "X \# fmset G as" by simp then have "X \ set (map (assocs G) as)" by (simp add: fmset_def) then have "\x. X = assocs G x \ x \ set as" by (induct as) auto then obtain x where X: "X = assocs G x" and xas: "x \ set as" by blast with ascarr have xcarr: "x \ carrier G" by blast from xas airr have xirr: "irreducible G x" by simp from xcarr and xirr and X show "\x. (x \ carrier G \ irreducible G x) \ X = assocs G x" by blast qed then obtain c cs where ccarr: "c \ carrier G" and cscarr: "set cs \ carrier G" and csirr: "wfactors G cs c" and csmset: "fmset G cs = fmset G as \# fmset G bs" by auto have "c gcdof a b" proof (simp add: isgcd_def, safe) from csmset have "fmset G cs \# fmset G as" by simp then show "c divides a" by (rule fmsubset_divides) fact+ next from csmset have "fmset G cs \# fmset G bs" by simp then show "c divides b" by (rule fmsubset_divides) fact+ next fix y assume "y \ carrier G" from wfactors_exist [OF this] obtain ys where yscarr: "set ys \ carrier G" and yfs: "wfactors G ys y" by blast assume "y divides a" then have ya: "fmset G ys \# fmset G as" by (rule divides_fmsubset) fact+ assume "y divides b" then have yb: "fmset G ys \# fmset G bs" by (rule divides_fmsubset) fact+ from ya yb csmset have "fmset G ys \# fmset G cs" by (simp add: subset_mset_def) then show "y divides c" by (rule fmsubset_divides) fact+ qed with ccarr show "\c. c \ carrier G \ c gcdof a b" by fast qed lemma (in factorial_monoid) lcmof_exists: assumes acarr: "a \ carrier G" and bcarr: "b \ carrier G" shows "\c. c \ carrier G \ c lcmof a b" proof - from wfactors_exist [OF acarr] obtain as where ascarr: "set as \ carrier G" and afs: "wfactors G as a" by blast from afs have airr: "\a \ set as. irreducible G a" by (fast elim: wfactorsE) from wfactors_exist [OF bcarr] obtain bs where bscarr: "set bs \ carrier G" and bfs: "wfactors G bs b" by blast from bfs have birr: "\b \ set bs. irreducible G b" by (fast elim: wfactorsE) have "\c cs. c \ carrier G \ set cs \ carrier G \ wfactors G cs c \ fmset G cs = (fmset G as - fmset G bs) + fmset G bs" proof (intro mset_wfactorsEx) fix X assume "X \# (fmset G as - fmset G bs) + fmset G bs" then have "X \# fmset G as \ X \# fmset G bs" by (auto dest: in_diffD) then consider "X \ set_mset (fmset G as)" | "X \ set_mset (fmset G bs)" by fast then show "\x. (x \ carrier G \ irreducible G x) \ X = assocs G x" proof cases case 1 then have "X \ set (map (assocs G) as)" by (simp add: fmset_def) then have "\x. x \ set as \ X = assocs G x" by (induct as) auto then obtain x where xas: "x \ set as" and X: "X = assocs G x" by auto with ascarr have xcarr: "x \ carrier G" by fast from xas airr have xirr: "irreducible G x" by simp from xcarr and xirr and X show ?thesis by fast next case 2 then have "X \ set (map (assocs G) bs)" by (simp add: fmset_def) then have "\x. x \ set bs \ X = assocs G x" by (induct as) auto then obtain x where xbs: "x \ set bs" and X: "X = assocs G x" by auto with bscarr have xcarr: "x \ carrier G" by fast from xbs birr have xirr: "irreducible G x" by simp from xcarr and xirr and X show ?thesis by fast qed qed then obtain c cs where ccarr: "c \ carrier G" and cscarr: "set cs \ carrier G" and csirr: "wfactors G cs c" and csmset: "fmset G cs = fmset G as - fmset G bs + fmset G bs" by auto have "c lcmof a b" proof (simp add: islcm_def, safe) from csmset have "fmset G as \# fmset G cs" by (simp add: subseteq_mset_def, force) then show "a divides c" by (rule fmsubset_divides) fact+ next from csmset have "fmset G bs \# fmset G cs" by (simp add: subset_mset_def) then show "b divides c" by (rule fmsubset_divides) fact+ next fix y assume "y \ carrier G" from wfactors_exist [OF this] obtain ys where yscarr: "set ys \ carrier G" and yfs: "wfactors G ys y" by blast assume "a divides y" then have ya: "fmset G as \# fmset G ys" by (rule divides_fmsubset) fact+ assume "b divides y" then have yb: "fmset G bs \# fmset G ys" by (rule divides_fmsubset) fact+ from ya yb csmset have "fmset G cs \# fmset G ys" using subset_eq_diff_conv subset_mset.le_diff_conv2 by fastforce then show "c divides y" by (rule fmsubset_divides) fact+ qed with ccarr show "\c. c \ carrier G \ c lcmof a b" by fast qed subsection \Conditions for Factoriality\ subsubsection \Gcd condition\ lemma (in gcd_condition_monoid) division_weak_lower_semilattice [simp]: "weak_lower_semilattice (division_rel G)" proof - interpret weak_partial_order "division_rel G" .. show ?thesis proof (unfold_locales, simp_all) fix x y assume carr: "x \ carrier G" "y \ carrier G" from gcdof_exists [OF this] obtain z where zcarr: "z \ carrier G" and isgcd: "z gcdof x y" by blast with carr have "greatest (division_rel G) z (Lower (division_rel G) {x, y})" by (subst gcdof_greatestLower[symmetric], simp+) then show "\z. greatest (division_rel G) z (Lower (division_rel G) {x, y})" by fast qed qed lemma (in gcd_condition_monoid) gcdof_cong_l: assumes "a' \ a" "a gcdof b c" "a' \ carrier G" and carr': "a \ carrier G" "b \ carrier G" "c \ carrier G" shows "a' gcdof b c" proof - interpret weak_lower_semilattice "division_rel G" by simp have "is_glb (division_rel G) a' {b, c}" by (subst greatest_Lower_cong_l[of _ a]) (simp_all add: assms gcdof_greatestLower[symmetric]) then have "a' \ carrier G \ a' gcdof b c" by (simp add: gcdof_greatestLower carr') then show ?thesis .. qed lemma (in gcd_condition_monoid) gcd_closed [simp]: assumes "a \ carrier G" "b \ carrier G" shows "somegcd G a b \ carrier G" proof - interpret weak_lower_semilattice "division_rel G" by simp show ?thesis using assms meet_closed by (simp add: somegcd_meet) qed lemma (in gcd_condition_monoid) gcd_isgcd: assumes "a \ carrier G" "b \ carrier G" shows "(somegcd G a b) gcdof a b" proof - interpret weak_lower_semilattice "division_rel G" by simp from assms have "somegcd G a b \ carrier G \ (somegcd G a b) gcdof a b" by (simp add: gcdof_greatestLower inf_of_two_greatest meet_def somegcd_meet) then show "(somegcd G a b) gcdof a b" by simp qed lemma (in gcd_condition_monoid) gcd_exists: assumes "a \ carrier G" "b \ carrier G" shows "\x\carrier G. x = somegcd G a b" proof - interpret weak_lower_semilattice "division_rel G" by simp show ?thesis by (metis assms gcd_closed) qed lemma (in gcd_condition_monoid) gcd_divides_l: assumes "a \ carrier G" "b \ carrier G" shows "(somegcd G a b) divides a" proof - interpret weak_lower_semilattice "division_rel G" by simp show ?thesis by (metis assms gcd_isgcd isgcd_def) qed lemma (in gcd_condition_monoid) gcd_divides_r: assumes "a \ carrier G" "b \ carrier G" shows "(somegcd G a b) divides b" proof - interpret weak_lower_semilattice "division_rel G" by simp show ?thesis by (metis assms gcd_isgcd isgcd_def) qed lemma (in gcd_condition_monoid) gcd_divides: assumes "z divides x" "z divides y" and L: "x \ carrier G" "y \ carrier G" "z \ carrier G" shows "z divides (somegcd G x y)" proof - interpret weak_lower_semilattice "division_rel G" by simp show ?thesis by (metis gcd_isgcd isgcd_def assms) qed lemma (in gcd_condition_monoid) gcd_cong_l: assumes "x \ x'" "x \ carrier G" "x' \ carrier G" "y \ carrier G" shows "somegcd G x y \ somegcd G x' y" proof - interpret weak_lower_semilattice "division_rel G" by simp show ?thesis using somegcd_meet assms by (metis eq_object.select_convs(1) meet_cong_l partial_object.select_convs(1)) qed lemma (in gcd_condition_monoid) gcd_cong_r: assumes "y \ y'" "x \ carrier G" "y \ carrier G" "y' \ carrier G" shows "somegcd G x y \ somegcd G x y'" proof - interpret weak_lower_semilattice "division_rel G" by simp show ?thesis by (meson associated_def assms gcd_closed gcd_divides gcd_divides_l gcd_divides_r monoid.divides_trans monoid_axioms) qed lemma (in gcd_condition_monoid) gcdI: assumes dvd: "a divides b" "a divides c" and others: "\y. \y\carrier G; y divides b; y divides c\ \ y divides a" and acarr: "a \ carrier G" and bcarr: "b \ carrier G" and ccarr: "c \ carrier G" shows "a \ somegcd G b c" proof - have "\a. a \ carrier G \ a gcdof b c" by (simp add: bcarr ccarr gcdof_exists) moreover have "\x. x \ carrier G \ x gcdof b c \ a \ x" by (simp add: acarr associated_def dvd isgcd_def others) ultimately show ?thesis unfolding somegcd_def by (blast intro: someI2_ex) qed lemma (in gcd_condition_monoid) gcdI2: assumes "a gcdof b c" and "a \ carrier G" and "b \ carrier G" and "c \ carrier G" shows "a \ somegcd G b c" using assms unfolding isgcd_def by (simp add: gcdI) lemma (in gcd_condition_monoid) SomeGcd_ex: assumes "finite A" "A \ carrier G" "A \ {}" shows "\x \ carrier G. x = SomeGcd G A" proof - interpret weak_lower_semilattice "division_rel G" by simp show ?thesis using finite_inf_closed by (simp add: assms SomeGcd_def) qed lemma (in gcd_condition_monoid) gcd_assoc: assumes "a \ carrier G" "b \ carrier G" "c \ carrier G" shows "somegcd G (somegcd G a b) c \ somegcd G a (somegcd G b c)" proof - interpret weak_lower_semilattice "division_rel G" by simp show ?thesis unfolding associated_def by (meson assms divides_trans gcd_divides gcd_divides_l gcd_divides_r gcd_exists) qed lemma (in gcd_condition_monoid) gcd_mult: assumes acarr: "a \ carrier G" and bcarr: "b \ carrier G" and ccarr: "c \ carrier G" shows "c \ somegcd G a b \ somegcd G (c \ a) (c \ b)" proof - (* following Jacobson, Basic Algebra, p.140 *) let ?d = "somegcd G a b" let ?e = "somegcd G (c \ a) (c \ b)" note carr[simp] = acarr bcarr ccarr have dcarr: "?d \ carrier G" by simp have ecarr: "?e \ carrier G" by simp note carr = carr dcarr ecarr have "?d divides a" by (simp add: gcd_divides_l) then have cd'ca: "c \ ?d divides (c \ a)" by (simp add: divides_mult_lI) have "?d divides b" by (simp add: gcd_divides_r) then have cd'cb: "c \ ?d divides (c \ b)" by (simp add: divides_mult_lI) from cd'ca cd'cb have cd'e: "c \ ?d divides ?e" by (rule gcd_divides) simp_all then obtain u where ucarr[simp]: "u \ carrier G" and e_cdu: "?e = c \ ?d \ u" by blast note carr = carr ucarr have "?e divides c \ a" by (rule gcd_divides_l) simp_all then obtain x where xcarr: "x \ carrier G" and ca_ex: "c \ a = ?e \ x" by blast with e_cdu have ca_cdux: "c \ a = c \ ?d \ u \ x" by simp from ca_cdux xcarr have "c \ a = c \ (?d \ u \ x)" by (simp add: m_assoc) then have "a = ?d \ u \ x" by (rule l_cancel[of c a]) (simp add: xcarr)+ then have du'a: "?d \ u divides a" by (rule dividesI[OF xcarr]) have "?e divides c \ b" by (intro gcd_divides_r) simp_all then obtain x where xcarr: "x \ carrier G" and cb_ex: "c \ b = ?e \ x" by blast with e_cdu have cb_cdux: "c \ b = c \ ?d \ u \ x" by simp from cb_cdux xcarr have "c \ b = c \ (?d \ u \ x)" by (simp add: m_assoc) with xcarr have "b = ?d \ u \ x" by (intro l_cancel[of c b]) simp_all then have du'b: "?d \ u divides b" by (intro dividesI[OF xcarr]) from du'a du'b carr have du'd: "?d \ u divides ?d" by (intro gcd_divides) simp_all then have uunit: "u \ Units G" proof (elim dividesE) fix v assume vcarr[simp]: "v \ carrier G" assume d: "?d = ?d \ u \ v" have "?d \ \ = ?d \ u \ v" by simp fact also have "?d \ u \ v = ?d \ (u \ v)" by (simp add: m_assoc) finally have "?d \ \ = ?d \ (u \ v)" . then have i2: "\ = u \ v" by (rule l_cancel) simp_all then have i1: "\ = v \ u" by (simp add: m_comm) from vcarr i1[symmetric] i2[symmetric] show "u \ Units G" by (auto simp: Units_def) qed from e_cdu uunit have "somegcd G (c \ a) (c \ b) \ c \ somegcd G a b" by (intro associatedI2[of u]) simp_all from this[symmetric] show "c \ somegcd G a b \ somegcd G (c \ a) (c \ b)" by simp qed lemma (in monoid) assoc_subst: assumes ab: "a \ b" and cP: "\a b. a \ carrier G \ b \ carrier G \ a \ b \ f a \ carrier G \ f b \ carrier G \ f a \ f b" and carr: "a \ carrier G" "b \ carrier G" shows "f a \ f b" using assms by auto lemma (in gcd_condition_monoid) relprime_mult: assumes abrelprime: "somegcd G a b \ \" and acrelprime: "somegcd G a c \ \" and carr[simp]: "a \ carrier G" "b \ carrier G" "c \ carrier G" shows "somegcd G a (b \ c) \ \" proof - have "c = c \ \" by simp also from abrelprime[symmetric] have "\ \ c \ somegcd G a b" by (rule assoc_subst) (simp add: mult_cong_r)+ also have "\ \ somegcd G (c \ a) (c \ b)" by (rule gcd_mult) fact+ finally have c: "c \ somegcd G (c \ a) (c \ b)" by simp from carr have a: "a \ somegcd G a (c \ a)" by (fast intro: gcdI divides_prod_l) have "somegcd G a (b \ c) \ somegcd G a (c \ b)" by (simp add: m_comm) also from a have "\ \ somegcd G (somegcd G a (c \ a)) (c \ b)" by (rule assoc_subst) (simp add: gcd_cong_l)+ also from gcd_assoc have "\ \ somegcd G a (somegcd G (c \ a) (c \ b))" by (rule assoc_subst) simp+ also from c[symmetric] have "\ \ somegcd G a c" by (rule assoc_subst) (simp add: gcd_cong_r)+ also note acrelprime finally show "somegcd G a (b \ c) \ \" by simp qed lemma (in gcd_condition_monoid) primeness_condition: "primeness_condition_monoid G" proof - have *: "p divides a \ p divides b" if pcarr[simp]: "p \ carrier G" and acarr[simp]: "a \ carrier G" and bcarr[simp]: "b \ carrier G" and pirr: "irreducible G p" and pdvdab: "p divides a \ b" for p a b proof - from pirr have pnunit: "p \ Units G" and r: "\b. \b \ carrier G; properfactor G b p\ \ b \ Units G" by (fast elim: irreducibleE)+ show "p divides a \ p divides b" proof (rule ccontr, clarsimp) assume npdvda: "\ p divides a" and npdvdb: "\ p divides b" have "\ \ somegcd G p a" proof (intro gcdI unit_divides) show "\y. \y \ carrier G; y divides p; y divides a\ \ y \ Units G" by (meson divides_trans npdvda pcarr properfactorI r) qed auto with pcarr acarr have pa: "somegcd G p a \ \" by (fast intro: associated_sym[of "\"] gcd_closed) have "\ \ somegcd G p b" proof (intro gcdI unit_divides) show "\y. \y \ carrier G; y divides p; y divides b\ \ y \ Units G" by (meson divides_trans npdvdb pcarr properfactorI r) qed auto with pcarr bcarr have pb: "somegcd G p b \ \" by (fast intro: associated_sym[of "\"] gcd_closed) have "p \ somegcd G p (a \ b)" using pdvdab by (simp add: gcdI2 isgcd_divides_l) also from pa pb pcarr acarr bcarr have "somegcd G p (a \ b) \ \" by (rule relprime_mult) finally have "p \ \" by simp with pcarr have "p \ Units G" by (fast intro: assoc_unit_l) with pnunit show False .. qed qed show ?thesis by unfold_locales (metis * primeI irreducibleE) qed sublocale gcd_condition_monoid \ primeness_condition_monoid by (rule primeness_condition) subsubsection \Divisor chain condition\ lemma (in divisor_chain_condition_monoid) wfactors_exist: assumes acarr: "a \ carrier G" shows "\as. set as \ carrier G \ wfactors G as a" proof - have r: "a \ carrier G \ (\as. set as \ carrier G \ wfactors G as a)" using division_wellfounded proof (induction rule: wf_induct_rule) case (less x) then have xcarr: "x \ carrier G" by auto show ?case proof (cases "x \ Units G") case True then show ?thesis by (metis bot.extremum list.set(1) unit_wfactors) next case xnunit: False show ?thesis proof (cases "irreducible G x") case True then show ?thesis by (rule_tac x="[x]" in exI) (simp add: wfactors_def xcarr) next case False then obtain y where ycarr: "y \ carrier G" and ynunit: "y \ Units G" and pfyx: "properfactor G y x" by (meson irreducible_def xnunit) obtain ys where yscarr: "set ys \ carrier G" and yfs: "wfactors G ys y" using less ycarr pfyx by blast then obtain z where zcarr: "z \ carrier G" and x: "x = y \ z" by (meson dividesE pfyx properfactorE2) from zcarr ycarr have "properfactor G z x" using m_comm properfactorI3 x ynunit by blast with less zcarr obtain zs where zscarr: "set zs \ carrier G" and zfs: "wfactors G zs z" by blast from yscarr zscarr have xscarr: "set (ys@zs) \ carrier G" by simp have "wfactors G (ys@zs) (y\z)" using xscarr ycarr yfs zcarr zfs by auto then have "wfactors G (ys@zs) x" by (simp add: x) with xscarr show "\xs. set xs \ carrier G \ wfactors G xs x" by fast qed qed qed from acarr show ?thesis by (rule r) qed subsubsection \Primeness condition\ lemma (in comm_monoid_cancel) multlist_prime_pos: assumes aprime: "prime G a" and carr: "a \ carrier G" and as: "set as \ carrier G" "a divides (foldr (\) as \)" shows "\i carrier G" "set as \ carrier G" and "a divides x \ foldr (\) as \" - by (auto simp: ) + by auto with carr aprime have "a divides x \ a divides foldr (\) as \" by (intro prime_divides) simp+ then show ?case using Cons.IH Cons.prems(1) by force qed proposition (in primeness_condition_monoid) wfactors_unique: assumes "wfactors G as a" "wfactors G as' a" and "a \ carrier G" "set as \ carrier G" "set as' \ carrier G" shows "essentially_equal G as as'" using assms proof (induct as arbitrary: a as') case Nil then have "a \ \" by (simp add: perm_wfactorsD) then have "as' = []" using Nil.prems assoc_unit_l unit_wfactors_empty by blast then show ?case by auto next case (Cons ah as) then have ahdvda: "ah divides a" using wfactors_dividesI by auto then obtain a' where a'carr: "a' \ carrier G" and a: "a = ah \ a'" by blast have carr_ah: "ah \ carrier G" "set as \ carrier G" using Cons.prems by fastforce+ have "ah \ foldr (\) as \ \ a" by (rule wfactorsE[OF \wfactors G (ah # as) a\]) auto then have "foldr (\) as \ \ a'" by (metis Cons.prems(4) a a'carr assoc_l_cancel insert_subset list.set(2) monoid.multlist_closed monoid_axioms) then have a'fs: "wfactors G as a'" by (meson Cons.prems(1) set_subset_Cons subset_iff wfactorsE wfactorsI) then have ahirr: "irreducible G ah" by (meson Cons.prems(1) list.set_intros(1) wfactorsE) with Cons have ahprime: "prime G ah" by (simp add: irreducible_prime) note ahdvda also have "a divides (foldr (\) as' \)" by (meson Cons.prems(2) associatedE wfactorsE) finally have "ah divides (foldr (\) as' \)" using Cons.prems(4) by auto with ahprime have "\i carrier G" and asi: "as'!i = ah \ x" by blast have irrasi: "irreducible G (as'!i)" using nth_mem[OF len] wfactorsE by (metis Cons.prems(2)) have asicarr[simp]: "as'!i \ carrier G" using len \set as' \ carrier G\ nth_mem by blast have asiah: "as'!i \ ah" by (metis \ah \ carrier G\ \x \ carrier G\ asi irrasi ahprime associatedI2 irreducible_prodE primeE) note setparts = set_take_subset[of i as'] set_drop_subset[of "Suc i" as'] have "\aa_1. aa_1 \ carrier G \ wfactors G (take i as') aa_1" using Cons by (metis setparts(1) subset_trans in_set_takeD wfactorsE wfactors_prod_exists) then obtain aa_1 where aa1carr [simp]: "aa_1 \ carrier G" and aa1fs: "wfactors G (take i as') aa_1" by auto obtain aa_2 where aa2carr [simp]: "aa_2 \ carrier G" and aa2fs: "wfactors G (drop (Suc i) as') aa_2" by (metis Cons.prems(2) Cons.prems(5) subset_code(1) in_set_dropD wfactors_def wfactors_prod_exists) have set_drop: "set (drop (Suc i) as') \ carrier G" using Cons.prems(5) setparts(2) by blast moreover have set_take: "set (take i as') \ carrier G" using Cons.prems(5) setparts by auto moreover have v1: "wfactors G (take i as' @ drop (Suc i) as') (aa_1 \ aa_2)" using aa1fs aa2fs \set as' \ carrier G\ by (force simp add: dest: in_set_takeD in_set_dropD) ultimately have v1': "wfactors G (as'!i # take i as' @ drop (Suc i) as') (as'!i \ (aa_1 \ aa_2))" using irrasi wfactors_mult_single by (simp add: irrasi v1 wfactors_mult_single) have "wfactors G (as'!i # drop (Suc i) as') (as'!i \ aa_2)" by (simp add: aa2fs irrasi set_drop wfactors_mult_single) with len aa1carr aa2carr aa1fs have v2: "wfactors G (take i as' @ as'!i # drop (Suc i) as') (aa_1 \ (as'!i \ aa_2))" using wfactors_mult by (simp add: set_take set_drop) from len have as': "as' = (take i as' @ as'!i # drop (Suc i) as')" by (simp add: Cons_nth_drop_Suc) have eer: "essentially_equal G (take i as' @ as'!i # drop (Suc i) as') as'" using Cons.prems(5) as' by auto with v2 aa1carr aa2carr nth_mem[OF len] have "aa_1 \ (as'!i \ aa_2) \ a" using Cons.prems as' comm_monoid_cancel.ee_wfactorsD is_comm_monoid_cancel by fastforce then have t1: "as'!i \ (aa_1 \ aa_2) \ a" by (metis aa1carr aa2carr asicarr m_lcomm) from asiah have "ah \ (aa_1 \ aa_2) \ as'!i \ (aa_1 \ aa_2)" by (simp add: \ah \ carrier G\ associated_sym mult_cong_l) also note t1 finally have "ah \ (aa_1 \ aa_2) \ a" using Cons.prems(3) carr_ah aa1carr aa2carr by blast with aa1carr aa2carr a'carr nth_mem[OF len] have a': "aa_1 \ aa_2 \ a'" using a assoc_l_cancel carr_ah(1) by blast note v1 also note a' finally have "wfactors G (take i as' @ drop (Suc i) as') a'" by (simp add: a'carr set_drop set_take) from a'fs this have "essentially_equal G as (take i as' @ drop (Suc i) as')" using Cons.hyps a'carr carr_ah(2) set_drop set_take by auto then obtain bs where \mset as = mset bs\ \bs [\] take i as' @ drop (Suc i) as'\ by (auto simp add: essentially_equal_def) with carr_ah have \mset (ah # as) = mset (ah # bs)\ \ah # bs [\] ah # take i as' @ drop (Suc i) as'\ by simp_all then have ee1: "essentially_equal G (ah # as) (ah # take i as' @ drop (Suc i) as')" unfolding essentially_equal_def by blast have ee2: "essentially_equal G (ah # take i as' @ drop (Suc i) as') (as' ! i # take i as' @ drop (Suc i) as')" proof (intro essentially_equalI) show "ah # take i as' @ drop (Suc i) as' <~~> ah # take i as' @ drop (Suc i) as'" by simp next show "ah # take i as' @ drop (Suc i) as' [\] as' ! i # take i as' @ drop (Suc i) as'" by (simp add: asiah associated_sym set_drop set_take) qed note ee1 also note ee2 also have "essentially_equal G (as' ! i # take i as' @ drop (Suc i) as') (take i as' @ as' ! i # drop (Suc i) as')" by (metis Cons.prems(5) as' essentially_equalI listassoc_refl perm_append_Cons) finally have "essentially_equal G (ah # as) (take i as' @ as' ! i # drop (Suc i) as')" using Cons.prems(4) set_drop set_take by auto then show ?case using as' by auto qed subsubsection \Application to factorial monoids\ text \Number of factors for wellfoundedness\ definition factorcount :: "_ \ 'a \ nat" where "factorcount G a = (THE c. \as. set as \ carrier G \ wfactors G as a \ c = length as)" lemma (in monoid) ee_length: assumes ee: "essentially_equal G as bs" shows "length as = length bs" by (rule essentially_equalE[OF ee]) (metis list_all2_conv_all_nth perm_length) lemma (in factorial_monoid) factorcount_exists: assumes carr[simp]: "a \ carrier G" shows "\c. \as. set as \ carrier G \ wfactors G as a \ c = length as" proof - have "\as. set as \ carrier G \ wfactors G as a" by (intro wfactors_exist) simp then obtain as where ascarr[simp]: "set as \ carrier G" and afs: "wfactors G as a" by (auto simp del: carr) have "\as'. set as' \ carrier G \ wfactors G as' a \ length as = length as'" by (metis afs ascarr assms ee_length wfactors_unique) then show "\c. \as'. set as' \ carrier G \ wfactors G as' a \ c = length as'" .. qed lemma (in factorial_monoid) factorcount_unique: assumes afs: "wfactors G as a" and acarr[simp]: "a \ carrier G" and ascarr: "set as \ carrier G" shows "factorcount G a = length as" proof - have "\ac. \as. set as \ carrier G \ wfactors G as a \ ac = length as" by (rule factorcount_exists) simp then obtain ac where alen: "\as. set as \ carrier G \ wfactors G as a \ ac = length as" by auto then have ac: "ac = factorcount G a" unfolding factorcount_def using ascarr by (blast intro: theI2 afs) from ascarr afs have "ac = length as" by (simp add: alen) with ac show ?thesis by simp qed lemma (in factorial_monoid) divides_fcount: assumes dvd: "a divides b" and acarr: "a \ carrier G" and bcarr:"b \ carrier G" shows "factorcount G a \ factorcount G b" proof (rule dividesE[OF dvd]) fix c from assms have "\as. set as \ carrier G \ wfactors G as a" by blast then obtain as where ascarr: "set as \ carrier G" and afs: "wfactors G as a" by blast with acarr have fca: "factorcount G a = length as" by (intro factorcount_unique) assume ccarr: "c \ carrier G" then have "\cs. set cs \ carrier G \ wfactors G cs c" by blast then obtain cs where cscarr: "set cs \ carrier G" and cfs: "wfactors G cs c" by blast note [simp] = acarr bcarr ccarr ascarr cscarr assume b: "b = a \ c" from afs cfs have "wfactors G (as@cs) (a \ c)" by (intro wfactors_mult) simp_all with b have "wfactors G (as@cs) b" by simp then have "factorcount G b = length (as@cs)" by (intro factorcount_unique) simp_all then have "factorcount G b = length as + length cs" by simp with fca show ?thesis by simp qed lemma (in factorial_monoid) associated_fcount: assumes acarr: "a \ carrier G" and bcarr: "b \ carrier G" and asc: "a \ b" shows "factorcount G a = factorcount G b" using assms by (auto simp: associated_def factorial_monoid.divides_fcount factorial_monoid_axioms le_antisym) lemma (in factorial_monoid) properfactor_fcount: assumes acarr: "a \ carrier G" and bcarr:"b \ carrier G" and pf: "properfactor G a b" shows "factorcount G a < factorcount G b" proof (rule properfactorE[OF pf], elim dividesE) fix c from assms have "\as. set as \ carrier G \ wfactors G as a" by blast then obtain as where ascarr: "set as \ carrier G" and afs: "wfactors G as a" by blast with acarr have fca: "factorcount G a = length as" by (intro factorcount_unique) assume ccarr: "c \ carrier G" then have "\cs. set cs \ carrier G \ wfactors G cs c" by blast then obtain cs where cscarr: "set cs \ carrier G" and cfs: "wfactors G cs c" by blast assume b: "b = a \ c" have "wfactors G (as@cs) (a \ c)" by (rule wfactors_mult) fact+ with b have "wfactors G (as@cs) b" by simp with ascarr cscarr bcarr have "factorcount G b = length (as@cs)" by (simp add: factorcount_unique) then have fcb: "factorcount G b = length as + length cs" by simp assume nbdvda: "\ b divides a" have "c \ Units G" proof assume cunit:"c \ Units G" have "b \ inv c = a \ c \ inv c" by (simp add: b) also from ccarr acarr cunit have "\ = a \ (c \ inv c)" by (fast intro: m_assoc) also from ccarr cunit have "\ = a \ \" by simp also from acarr have "\ = a" by simp finally have "a = b \ inv c" by simp with ccarr cunit have "b divides a" by (fast intro: dividesI[of "inv c"]) with nbdvda show False by simp qed with cfs have "length cs > 0" by (metis Units_one_closed assoc_unit_r ccarr foldr.simps(1) id_apply length_greater_0_conv wfactors_def) with fca fcb show ?thesis by simp qed sublocale factorial_monoid \ divisor_chain_condition_monoid apply unfold_locales apply (rule wfUNIVI) apply (rule measure_induct[of "factorcount G"]) using properfactor_fcount by auto sublocale factorial_monoid \ primeness_condition_monoid by standard (rule irreducible_prime) lemma (in factorial_monoid) primeness_condition: "primeness_condition_monoid G" .. lemma (in factorial_monoid) gcd_condition [simp]: "gcd_condition_monoid G" by standard (rule gcdof_exists) sublocale factorial_monoid \ gcd_condition_monoid by standard (rule gcdof_exists) lemma (in factorial_monoid) division_weak_lattice [simp]: "weak_lattice (division_rel G)" proof - interpret weak_lower_semilattice "division_rel G" by simp show "weak_lattice (division_rel G)" proof (unfold_locales, simp_all) fix x y assume carr: "x \ carrier G" "y \ carrier G" from lcmof_exists [OF this] obtain z where zcarr: "z \ carrier G" and isgcd: "z lcmof x y" by blast with carr have "least (division_rel G) z (Upper (division_rel G) {x, y})" by (simp add: lcmof_leastUpper[symmetric]) then show "\z. least (division_rel G) z (Upper (division_rel G) {x, y})" by blast qed qed subsection \Factoriality Theorems\ theorem factorial_condition_one: (* Jacobson theorem 2.21 *) "divisor_chain_condition_monoid G \ primeness_condition_monoid G \ factorial_monoid G" proof (rule iffI, clarify) assume dcc: "divisor_chain_condition_monoid G" and pc: "primeness_condition_monoid G" interpret divisor_chain_condition_monoid "G" by (rule dcc) interpret primeness_condition_monoid "G" by (rule pc) show "factorial_monoid G" by (fast intro: factorial_monoidI wfactors_exist wfactors_unique) next assume "factorial_monoid G" then interpret factorial_monoid "G" . show "divisor_chain_condition_monoid G \ primeness_condition_monoid G" by rule unfold_locales qed theorem factorial_condition_two: (* Jacobson theorem 2.22 *) "divisor_chain_condition_monoid G \ gcd_condition_monoid G \ factorial_monoid G" proof (rule iffI, clarify) assume dcc: "divisor_chain_condition_monoid G" and gc: "gcd_condition_monoid G" interpret divisor_chain_condition_monoid "G" by (rule dcc) interpret gcd_condition_monoid "G" by (rule gc) show "factorial_monoid G" by (simp add: factorial_condition_one[symmetric], rule, unfold_locales) next assume "factorial_monoid G" then interpret factorial_monoid "G" . show "divisor_chain_condition_monoid G \ gcd_condition_monoid G" by rule unfold_locales qed end diff --git a/src/HOL/Analysis/Abstract_Topology.thy b/src/HOL/Analysis/Abstract_Topology.thy --- a/src/HOL/Analysis/Abstract_Topology.thy +++ b/src/HOL/Analysis/Abstract_Topology.thy @@ -1,4621 +1,4621 @@ (* Author: L C Paulson, University of Cambridge [ported from HOL Light] *) section \Operators involving abstract topology\ theory Abstract_Topology imports Complex_Main "HOL-Library.Set_Idioms" "HOL-Library.FuncSet" begin subsection \General notion of a topology as a value\ definition\<^marker>\tag important\ istopology :: "('a set \ bool) \ bool" where "istopology L \ (\S T. L S \ L T \ L (S \ T)) \ (\\. (\K\\. L K) \ L (\\))" typedef\<^marker>\tag important\ 'a topology = "{L::('a set) \ bool. istopology L}" morphisms "openin" "topology" unfolding istopology_def by blast lemma istopology_openin[iff]: "istopology(openin U)" using openin[of U] by blast lemma istopology_open[iff]: "istopology open" by (auto simp: istopology_def) lemma topology_inverse' [simp]: "istopology U \ openin (topology U) = U" using topology_inverse[unfolded mem_Collect_eq] . lemma topology_inverse_iff: "istopology U \ openin (topology U) = U" by (metis istopology_openin topology_inverse') lemma topology_eq: "T1 = T2 \ (\S. openin T1 S \ openin T2 S)" proof assume "T1 = T2" then show "\S. openin T1 S \ openin T2 S" by simp next assume H: "\S. openin T1 S \ openin T2 S" then have "openin T1 = openin T2" by (simp add: fun_eq_iff) then have "topology (openin T1) = topology (openin T2)" by simp then show "T1 = T2" unfolding openin_inverse . qed text\The "universe": the union of all sets in the topology.\ definition "topspace T = \{S. openin T S}" subsubsection \Main properties of open sets\ proposition openin_clauses: fixes U :: "'a topology" shows "openin U {}" "\S T. openin U S \ openin U T \ openin U (S\T)" "\K. (\S \ K. openin U S) \ openin U (\K)" using openin[of U] unfolding istopology_def by auto lemma openin_subset: "openin U S \ S \ topspace U" unfolding topspace_def by blast lemma openin_empty[simp]: "openin U {}" by (rule openin_clauses) lemma openin_Int[intro]: "openin U S \ openin U T \ openin U (S \ T)" by (rule openin_clauses) lemma openin_Union[intro]: "(\S. S \ K \ openin U S) \ openin U (\K)" using openin_clauses by blast lemma openin_Un[intro]: "openin U S \ openin U T \ openin U (S \ T)" using openin_Union[of "{S,T}" U] by auto lemma openin_topspace[intro, simp]: "openin U (topspace U)" by (force simp: openin_Union topspace_def) lemma openin_subopen: "openin U S \ (\x \ S. \T. openin U T \ x \ T \ T \ S)" (is "?lhs \ ?rhs") proof assume ?lhs then show ?rhs by auto next assume H: ?rhs let ?t = "\{T. openin U T \ T \ S}" have "openin U ?t" by (force simp: openin_Union) also have "?t = S" using H by auto finally show "openin U S" . qed lemma openin_INT [intro]: assumes "finite I" "\i. i \ I \ openin T (U i)" shows "openin T ((\i \ I. U i) \ topspace T)" using assms by (induct, auto simp: inf_sup_aci(2) openin_Int) lemma openin_INT2 [intro]: assumes "finite I" "I \ {}" "\i. i \ I \ openin T (U i)" shows "openin T (\i \ I. U i)" proof - have "(\i \ I. U i) \ topspace T" using \I \ {}\ openin_subset[OF assms(3)] by auto then show ?thesis using openin_INT[of _ _ U, OF assms(1) assms(3)] by (simp add: inf.absorb2 inf_commute) qed lemma openin_Inter [intro]: assumes "finite \" "\ \ {}" "\X. X \ \ \ openin T X" shows "openin T (\\)" by (metis (full_types) assms openin_INT2 image_ident) lemma openin_Int_Inter: assumes "finite \" "openin T U" "\X. X \ \ \ openin T X" shows "openin T (U \ \\)" using openin_Inter [of "insert U \"] assms by auto subsubsection \Closed sets\ definition\<^marker>\tag important\ closedin :: "'a topology \ 'a set \ bool" where "closedin U S \ S \ topspace U \ openin U (topspace U - S)" lemma closedin_subset: "closedin U S \ S \ topspace U" by (metis closedin_def) lemma closedin_empty[simp]: "closedin U {}" by (simp add: closedin_def) lemma closedin_topspace[intro, simp]: "closedin U (topspace U)" by (simp add: closedin_def) lemma closedin_Un[intro]: "closedin U S \ closedin U T \ closedin U (S \ T)" by (auto simp: Diff_Un closedin_def) lemma Diff_Inter[intro]: "A - \S = \{A - s|s. s\S}" by auto lemma closedin_Union: assumes "finite S" "\T. T \ S \ closedin U T" shows "closedin U (\S)" using assms by induction auto lemma closedin_Inter[intro]: assumes Ke: "K \ {}" and Kc: "\S. S \K \ closedin U S" shows "closedin U (\K)" using Ke Kc unfolding closedin_def Diff_Inter by auto lemma closedin_INT[intro]: assumes "A \ {}" "\x. x \ A \ closedin U (B x)" shows "closedin U (\x\A. B x)" using assms by blast lemma closedin_Int[intro]: "closedin U S \ closedin U T \ closedin U (S \ T)" using closedin_Inter[of "{S,T}" U] by auto lemma openin_closedin_eq: "openin U S \ S \ topspace U \ closedin U (topspace U - S)" by (metis Diff_subset closedin_def double_diff equalityD1 openin_subset) lemma topology_finer_closedin: "topspace X = topspace Y \ (\S. openin Y S \ openin X S) \ (\S. closedin Y S \ closedin X S)" by (metis closedin_def openin_closedin_eq) lemma openin_closedin: "S \ topspace U \ (openin U S \ closedin U (topspace U - S))" by (simp add: openin_closedin_eq) lemma openin_diff[intro]: assumes oS: "openin U S" and cT: "closedin U T" shows "openin U (S - T)" proof - have "S - T = S \ (topspace U - T)" using openin_subset[of U S] oS cT by (auto simp: topspace_def openin_subset) then show ?thesis using oS cT by (auto simp: closedin_def) qed lemma closedin_diff[intro]: assumes oS: "closedin U S" and cT: "openin U T" shows "closedin U (S - T)" proof - have "S - T = S \ (topspace U - T)" using closedin_subset[of U S] oS cT by (auto simp: topspace_def) then show ?thesis using oS cT by (auto simp: openin_closedin_eq) qed subsection\The discrete topology\ definition discrete_topology where "discrete_topology U \ topology (\S. S \ U)" lemma openin_discrete_topology [simp]: "openin (discrete_topology U) S \ S \ U" proof - have "istopology (\S. S \ U)" by (auto simp: istopology_def) then show ?thesis by (simp add: discrete_topology_def topology_inverse') qed lemma topspace_discrete_topology [simp]: "topspace(discrete_topology U) = U" by (meson openin_discrete_topology openin_subset openin_topspace order_refl subset_antisym) lemma closedin_discrete_topology [simp]: "closedin (discrete_topology U) S \ S \ U" by (simp add: closedin_def) lemma discrete_topology_unique: "discrete_topology U = X \ topspace X = U \ (\x \ U. openin X {x})" (is "?lhs = ?rhs") proof assume R: ?rhs then have "openin X S" if "S \ U" for S using openin_subopen subsetD that by fastforce moreover have "x \ topspace X" if "openin X S" and "x \ S" for x S using openin_subset that by blast ultimately show ?lhs using R by (auto simp: topology_eq) qed auto lemma discrete_topology_unique_alt: "discrete_topology U = X \ topspace X \ U \ (\x \ U. openin X {x})" using openin_subset by (auto simp: discrete_topology_unique) lemma subtopology_eq_discrete_topology_empty: "X = discrete_topology {} \ topspace X = {}" using discrete_topology_unique [of "{}" X] by auto lemma subtopology_eq_discrete_topology_sing: "X = discrete_topology {a} \ topspace X = {a}" by (metis discrete_topology_unique openin_topspace singletonD) subsection \Subspace topology\ definition\<^marker>\tag important\ subtopology :: "'a topology \ 'a set \ 'a topology" where "subtopology U V = topology (\T. \S. T = S \ V \ openin U S)" lemma istopology_subtopology: "istopology (\T. \S. T = S \ V \ openin U S)" (is "istopology ?L") proof - have "?L {}" by blast { fix A B assume A: "?L A" and B: "?L B" from A B obtain Sa and Sb where Sa: "openin U Sa" "A = Sa \ V" and Sb: "openin U Sb" "B = Sb \ V" by blast have "A \ B = (Sa \ Sb) \ V" "openin U (Sa \ Sb)" using Sa Sb by blast+ then have "?L (A \ B)" by blast } moreover { fix K assume K: "K \ Collect ?L" have th0: "Collect ?L = (\S. S \ V) ` Collect (openin U)" by blast from K[unfolded th0 subset_image_iff] obtain Sk where Sk: "Sk \ Collect (openin U)" "K = (\S. S \ V) ` Sk" by blast have "\K = (\Sk) \ V" using Sk by auto moreover have "openin U (\Sk)" using Sk by (auto simp: subset_eq) ultimately have "?L (\K)" by blast } ultimately show ?thesis unfolding subset_eq mem_Collect_eq istopology_def by auto qed lemma openin_subtopology: "openin (subtopology U V) S \ (\T. openin U T \ S = T \ V)" unfolding subtopology_def topology_inverse'[OF istopology_subtopology] by auto lemma openin_subtopology_Int: "openin X S \ openin (subtopology X T) (S \ T)" using openin_subtopology by auto lemma openin_subtopology_Int2: "openin X T \ openin (subtopology X S) (S \ T)" using openin_subtopology by auto lemma openin_subtopology_diff_closed: "\S \ topspace X; closedin X T\ \ openin (subtopology X S) (S - T)" unfolding closedin_def openin_subtopology by (rule_tac x="topspace X - T" in exI) auto lemma openin_relative_to: "(openin X relative_to S) = openin (subtopology X S)" by (force simp: relative_to_def openin_subtopology) lemma topspace_subtopology [simp]: "topspace (subtopology U V) = topspace U \ V" by (auto simp: topspace_def openin_subtopology) lemma topspace_subtopology_subset: "S \ topspace X \ topspace(subtopology X S) = S" by (simp add: inf.absorb_iff2) lemma closedin_subtopology: "closedin (subtopology U V) S \ (\T. closedin U T \ S = T \ V)" unfolding closedin_def topspace_subtopology by (auto simp: openin_subtopology) lemma openin_subtopology_refl: "openin (subtopology U V) V \ V \ topspace U" unfolding openin_subtopology by auto (metis IntD1 in_mono openin_subset) lemma subtopology_subtopology: "subtopology (subtopology X S) T = subtopology X (S \ T)" proof - have eq: "\T'. (\S'. T' = S' \ T \ (\T. openin X T \ S' = T \ S)) = (\Sa. T' = Sa \ (S \ T) \ openin X Sa)" by (metis inf_assoc) have "subtopology (subtopology X S) T = topology (\Ta. \Sa. Ta = Sa \ T \ openin (subtopology X S) Sa)" by (simp add: subtopology_def) also have "\ = subtopology X (S \ T)" by (simp add: openin_subtopology eq) (simp add: subtopology_def) finally show ?thesis . qed lemma openin_subtopology_alt: "openin (subtopology X U) S \ S \ (\T. U \ T) ` Collect (openin X)" by (simp add: image_iff inf_commute openin_subtopology) lemma closedin_subtopology_alt: "closedin (subtopology X U) S \ S \ (\T. U \ T) ` Collect (closedin X)" by (simp add: image_iff inf_commute closedin_subtopology) lemma subtopology_superset: assumes UV: "topspace U \ V" shows "subtopology U V = U" proof - { fix S { fix T assume T: "openin U T" "S = T \ V" from T openin_subset[OF T(1)] UV have eq: "S = T" by blast have "openin U S" unfolding eq using T by blast } moreover { assume S: "openin U S" then have "\T. openin U T \ S = T \ V" using openin_subset[OF S] UV by auto } ultimately have "(\T. openin U T \ S = T \ V) \ openin U S" by blast } then show ?thesis unfolding topology_eq openin_subtopology by blast qed lemma subtopology_topspace[simp]: "subtopology U (topspace U) = U" by (simp add: subtopology_superset) lemma subtopology_UNIV[simp]: "subtopology U UNIV = U" by (simp add: subtopology_superset) lemma subtopology_restrict: "subtopology X (topspace X \ S) = subtopology X S" by (metis subtopology_subtopology subtopology_topspace) lemma openin_subtopology_empty: "openin (subtopology U {}) S \ S = {}" by (metis Int_empty_right openin_empty openin_subtopology) lemma closedin_subtopology_empty: "closedin (subtopology U {}) S \ S = {}" by (metis Int_empty_right closedin_empty closedin_subtopology) lemma closedin_subtopology_refl [simp]: "closedin (subtopology U X) X \ X \ topspace U" by (metis closedin_def closedin_topspace inf.absorb_iff2 le_inf_iff topspace_subtopology) lemma closedin_topspace_empty: "topspace T = {} \ (closedin T S \ S = {})" by (simp add: closedin_def) lemma open_in_topspace_empty: "topspace X = {} \ openin X S \ S = {}" by (simp add: openin_closedin_eq) lemma openin_imp_subset: "openin (subtopology U S) T \ T \ S" by (metis Int_iff openin_subtopology subsetI) lemma closedin_imp_subset: "closedin (subtopology U S) T \ T \ S" by (simp add: closedin_def) lemma openin_open_subtopology: "openin X S \ openin (subtopology X S) T \ openin X T \ T \ S" by (metis inf.orderE openin_Int openin_imp_subset openin_subtopology) lemma closedin_closed_subtopology: "closedin X S \ (closedin (subtopology X S) T \ closedin X T \ T \ S)" by (metis closedin_Int closedin_imp_subset closedin_subtopology inf.orderE) lemma openin_subtopology_Un: "\openin (subtopology X T) S; openin (subtopology X U) S\ \ openin (subtopology X (T \ U)) S" by (simp add: openin_subtopology) blast lemma closedin_subtopology_Un: "\closedin (subtopology X T) S; closedin (subtopology X U) S\ \ closedin (subtopology X (T \ U)) S" by (simp add: closedin_subtopology) blast lemma openin_trans_full: "\openin (subtopology X U) S; openin X U\ \ openin X S" by (simp add: openin_open_subtopology) subsection \The canonical topology from the underlying type class\ abbreviation\<^marker>\tag important\ euclidean :: "'a::topological_space topology" where "euclidean \ topology open" abbreviation top_of_set :: "'a::topological_space set \ 'a topology" where "top_of_set \ subtopology (topology open)" lemma open_openin: "open S \ openin euclidean S" by (simp add: istopology_open topology_inverse') declare open_openin [symmetric, simp] lemma topspace_euclidean [simp]: "topspace euclidean = UNIV" by (force simp: topspace_def) lemma topspace_euclidean_subtopology[simp]: "topspace (top_of_set S) = S" by (simp) lemma closed_closedin: "closed S \ closedin euclidean S" by (simp add: closed_def closedin_def Compl_eq_Diff_UNIV) declare closed_closedin [symmetric, simp] lemma openin_subtopology_self [simp]: "openin (top_of_set S) S" by (metis openin_topspace topspace_euclidean_subtopology) subsubsection\The most basic facts about the usual topology and metric on R\ abbreviation euclideanreal :: "real topology" where "euclideanreal \ topology open" subsection \Basic "localization" results are handy for connectedness.\ lemma openin_open: "openin (top_of_set U) S \ (\T. open T \ (S = U \ T))" by (auto simp: openin_subtopology) lemma openin_Int_open: "\openin (top_of_set U) S; open T\ \ openin (top_of_set U) (S \ T)" by (metis open_Int Int_assoc openin_open) lemma openin_open_Int[intro]: "open S \ openin (top_of_set U) (U \ S)" by (auto simp: openin_open) lemma open_openin_trans[trans]: "open S \ open T \ T \ S \ openin (top_of_set S) T" by (metis Int_absorb1 openin_open_Int) lemma open_subset: "S \ T \ open S \ openin (top_of_set T) S" by (auto simp: openin_open) lemma closedin_closed: "closedin (top_of_set U) S \ (\T. closed T \ S = U \ T)" by (simp add: closedin_subtopology Int_ac) lemma closedin_closed_Int: "closed S \ closedin (top_of_set U) (U \ S)" by (metis closedin_closed) lemma closed_subset: "S \ T \ closed S \ closedin (top_of_set T) S" by (auto simp: closedin_closed) lemma closedin_closed_subset: "\closedin (top_of_set U) V; T \ U; S = V \ T\ \ closedin (top_of_set T) S" by (metis (no_types, lifting) Int_assoc Int_commute closedin_closed inf.orderE) lemma finite_imp_closedin: fixes S :: "'a::t1_space set" shows "\finite S; S \ T\ \ closedin (top_of_set T) S" by (simp add: finite_imp_closed closed_subset) lemma closedin_singleton [simp]: fixes a :: "'a::t1_space" shows "closedin (top_of_set U) {a} \ a \ U" using closedin_subset by (force intro: closed_subset) lemma openin_euclidean_subtopology_iff: fixes S U :: "'a::metric_space set" shows "openin (top_of_set U) S \ S \ U \ (\x\S. \e>0. \x'\U. dist x' x < e \ x'\ S)" (is "?lhs \ ?rhs") proof assume ?lhs then show ?rhs unfolding openin_open open_dist by blast next define T where "T = {x. \a\S. \d>0. (\y\U. dist y a < d \ y \ S) \ dist x a < d}" have 1: "\x\T. \e>0. \y. dist y x < e \ y \ T" unfolding T_def apply clarsimp apply (rule_tac x="d - dist x a" in exI) by (metis add_0_left dist_commute dist_triangle_lt less_diff_eq) assume ?rhs then have 2: "S = U \ T" unfolding T_def by auto (metis dist_self) from 1 2 show ?lhs unfolding openin_open open_dist by fast qed lemma connected_openin: "connected S \ \(\E1 E2. openin (top_of_set S) E1 \ openin (top_of_set S) E2 \ S \ E1 \ E2 \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" unfolding connected_def openin_open disjoint_iff_not_equal by blast lemma connected_openin_eq: "connected S \ \(\E1 E2. openin (top_of_set S) E1 \ openin (top_of_set S) E2 \ E1 \ E2 = S \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" unfolding connected_openin by (metis (no_types, lifting) Un_subset_iff openin_imp_subset subset_antisym) lemma connected_closedin: "connected S \ (\E1 E2. closedin (top_of_set S) E1 \ closedin (top_of_set S) E2 \ S \ E1 \ E2 \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (auto simp add: connected_closed closedin_closed) next assume R: ?rhs then show ?lhs proof (clarsimp simp add: connected_closed closedin_closed) fix A B assume s_sub: "S \ A \ B" "B \ S \ {}" and disj: "A \ B \ S = {}" and cl: "closed A" "closed B" have "S \ (A \ B) = S" using s_sub(1) by auto have "S - A = B \ S" using Diff_subset_conv Un_Diff_Int disj s_sub(1) by auto then have "S \ A = {}" by (metis Diff_Diff_Int Diff_disjoint Un_Diff_Int R cl closedin_closed_Int inf_commute order_refl s_sub(2)) then show "A \ S = {}" by blast qed qed lemma connected_closedin_eq: "connected S \ \(\E1 E2. closedin (top_of_set S) E1 \ closedin (top_of_set S) E2 \ E1 \ E2 = S \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" unfolding connected_closedin by (metis Un_subset_iff closedin_imp_subset subset_antisym) text \These "transitivity" results are handy too\ lemma openin_trans[trans]: "openin (top_of_set T) S \ openin (top_of_set U) T \ openin (top_of_set U) S" by (metis openin_Int_open openin_open) lemma openin_open_trans: "openin (top_of_set T) S \ open T \ open S" by (auto simp: openin_open intro: openin_trans) lemma closedin_trans[trans]: "closedin (top_of_set T) S \ closedin (top_of_set U) T \ closedin (top_of_set U) S" by (auto simp: closedin_closed closed_Inter Int_assoc) lemma closedin_closed_trans: "closedin (top_of_set T) S \ closed T \ closed S" by (auto simp: closedin_closed intro: closedin_trans) lemma openin_subtopology_Int_subset: "\openin (top_of_set u) (u \ S); v \ u\ \ openin (top_of_set v) (v \ S)" by (auto simp: openin_subtopology) lemma openin_open_eq: "open s \ (openin (top_of_set s) t \ open t \ t \ s)" using open_subset openin_open_trans openin_subset by fastforce subsection\Derived set (set of limit points)\ definition derived_set_of :: "'a topology \ 'a set \ 'a set" (infixl "derived'_set'_of" 80) where "X derived_set_of S \ {x \ topspace X. (\T. x \ T \ openin X T \ (\y\x. y \ S \ y \ T))}" lemma derived_set_of_restrict [simp]: "X derived_set_of (topspace X \ S) = X derived_set_of S" by (simp add: derived_set_of_def) (metis openin_subset subset_iff) lemma in_derived_set_of: "x \ X derived_set_of S \ x \ topspace X \ (\T. x \ T \ openin X T \ (\y\x. y \ S \ y \ T))" by (simp add: derived_set_of_def) lemma derived_set_of_subset_topspace: "X derived_set_of S \ topspace X" by (auto simp add: derived_set_of_def) lemma derived_set_of_subtopology: "(subtopology X U) derived_set_of S = U \ (X derived_set_of (U \ S))" by (simp add: derived_set_of_def openin_subtopology) blast lemma derived_set_of_subset_subtopology: "(subtopology X S) derived_set_of T \ S" by (simp add: derived_set_of_subtopology) lemma derived_set_of_empty [simp]: "X derived_set_of {} = {}" by (auto simp: derived_set_of_def) lemma derived_set_of_mono: "S \ T \ X derived_set_of S \ X derived_set_of T" unfolding derived_set_of_def by blast lemma derived_set_of_Un: "X derived_set_of (S \ T) = X derived_set_of S \ X derived_set_of T" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (clarsimp simp: in_derived_set_of) (metis IntE IntI openin_Int) show "?rhs \ ?lhs" by (simp add: derived_set_of_mono) qed lemma derived_set_of_Union: "finite \ \ X derived_set_of (\\) = (\S \ \. X derived_set_of S)" proof (induction \ rule: finite_induct) case (insert S \) then show ?case by (simp add: derived_set_of_Un) qed auto lemma derived_set_of_topspace: "X derived_set_of (topspace X) = {x \ topspace X. \ openin X {x}}" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (auto simp: in_derived_set_of) show "?rhs \ ?lhs" by (clarsimp simp: in_derived_set_of) (metis openin_closedin_eq openin_subopen singletonD subset_eq) qed lemma discrete_topology_unique_derived_set: "discrete_topology U = X \ topspace X = U \ X derived_set_of U = {}" by (auto simp: discrete_topology_unique derived_set_of_topspace) lemma subtopology_eq_discrete_topology_eq: "subtopology X U = discrete_topology U \ U \ topspace X \ U \ X derived_set_of U = {}" using discrete_topology_unique_derived_set [of U "subtopology X U"] by (auto simp: eq_commute derived_set_of_subtopology) lemma subtopology_eq_discrete_topology: "S \ topspace X \ S \ X derived_set_of S = {} \ subtopology X S = discrete_topology S" by (simp add: subtopology_eq_discrete_topology_eq) lemma subtopology_eq_discrete_topology_gen: "S \ X derived_set_of S = {} \ subtopology X S = discrete_topology(topspace X \ S)" by (metis Int_lower1 derived_set_of_restrict inf_assoc inf_bot_right subtopology_eq_discrete_topology_eq subtopology_subtopology subtopology_topspace) lemma subtopology_discrete_topology [simp]: "subtopology (discrete_topology U) S = discrete_topology(U \ S)" proof - have "(\T. \Sa. T = Sa \ S \ Sa \ U) = (\Sa. Sa \ U \ Sa \ S)" by force then show ?thesis by (simp add: subtopology_def) (simp add: discrete_topology_def) qed lemma openin_Int_derived_set_of_subset: "openin X S \ S \ X derived_set_of T \ X derived_set_of (S \ T)" by (auto simp: derived_set_of_def) lemma openin_Int_derived_set_of_eq: assumes "openin X S" shows "S \ X derived_set_of T = S \ X derived_set_of (S \ T)" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (simp add: assms openin_Int_derived_set_of_subset) show "?rhs \ ?lhs" by (metis derived_set_of_mono inf_commute inf_le1 inf_mono order_refl) qed subsection\ Closure with respect to a topological space\ definition closure_of :: "'a topology \ 'a set \ 'a set" (infixr "closure'_of" 80) where "X closure_of S \ {x \ topspace X. \T. x \ T \ openin X T \ (\y \ S. y \ T)}" lemma closure_of_restrict: "X closure_of S = X closure_of (topspace X \ S)" unfolding closure_of_def using openin_subset by blast lemma in_closure_of: "x \ X closure_of S \ x \ topspace X \ (\T. x \ T \ openin X T \ (\y. y \ S \ y \ T))" by (auto simp: closure_of_def) lemma closure_of: "X closure_of S = topspace X \ (S \ X derived_set_of S)" by (fastforce simp: in_closure_of in_derived_set_of) lemma closure_of_alt: "X closure_of S = topspace X \ S \ X derived_set_of S" using derived_set_of_subset_topspace [of X S] unfolding closure_of_def in_derived_set_of by safe (auto simp: in_derived_set_of) lemma derived_set_of_subset_closure_of: "X derived_set_of S \ X closure_of S" by (fastforce simp: closure_of_def in_derived_set_of) lemma closure_of_subtopology: "(subtopology X U) closure_of S = U \ (X closure_of (U \ S))" unfolding closure_of_def topspace_subtopology openin_subtopology by safe (metis (full_types) IntI Int_iff inf.commute)+ lemma closure_of_empty [simp]: "X closure_of {} = {}" by (simp add: closure_of_alt) lemma closure_of_topspace [simp]: "X closure_of topspace X = topspace X" by (simp add: closure_of) lemma closure_of_UNIV [simp]: "X closure_of UNIV = topspace X" by (simp add: closure_of) lemma closure_of_subset_topspace: "X closure_of S \ topspace X" by (simp add: closure_of) lemma closure_of_subset_subtopology: "(subtopology X S) closure_of T \ S" by (simp add: closure_of_subtopology) lemma closure_of_mono: "S \ T \ X closure_of S \ X closure_of T" by (fastforce simp add: closure_of_def) lemma closure_of_subtopology_subset: "(subtopology X U) closure_of S \ (X closure_of S)" unfolding closure_of_subtopology by clarsimp (meson closure_of_mono contra_subsetD inf.cobounded2) lemma closure_of_subtopology_mono: "T \ U \ (subtopology X T) closure_of S \ (subtopology X U) closure_of S" unfolding closure_of_subtopology by auto (meson closure_of_mono inf_mono subset_iff) lemma closure_of_Un [simp]: "X closure_of (S \ T) = X closure_of S \ X closure_of T" by (simp add: Un_assoc Un_left_commute closure_of_alt derived_set_of_Un inf_sup_distrib1) lemma closure_of_Union: "finite \ \ X closure_of (\\) = (\S \ \. X closure_of S)" by (induction \ rule: finite_induct) auto lemma closure_of_subset: "S \ topspace X \ S \ X closure_of S" by (auto simp: closure_of_def) lemma closure_of_subset_Int: "topspace X \ S \ X closure_of S" by (auto simp: closure_of_def) lemma closure_of_subset_eq: "S \ topspace X \ X closure_of S \ S \ closedin X S" proof - have "openin X (topspace X - S)" if "\x. \x \ topspace X; \T. x \ T \ openin X T \ S \ T \ {}\ \ x \ S" apply (subst openin_subopen) by (metis Diff_iff Diff_mono Diff_triv inf.commute openin_subset order_refl that) then show ?thesis by (auto simp: closedin_def closure_of_def disjoint_iff_not_equal) qed lemma closure_of_eq: "X closure_of S = S \ closedin X S" proof (cases "S \ topspace X") case True then show ?thesis by (metis closure_of_subset closure_of_subset_eq set_eq_subset) next case False then show ?thesis using closure_of closure_of_subset_eq by fastforce qed lemma closedin_contains_derived_set: "closedin X S \ X derived_set_of S \ S \ S \ topspace X" proof (intro iffI conjI) show "closedin X S \ X derived_set_of S \ S" using closure_of_eq derived_set_of_subset_closure_of by fastforce show "closedin X S \ S \ topspace X" using closedin_subset by blast show "X derived_set_of S \ S \ S \ topspace X \ closedin X S" by (metis closure_of closure_of_eq inf.absorb_iff2 sup.orderE) qed lemma derived_set_subset_gen: "X derived_set_of S \ S \ closedin X (topspace X \ S)" by (simp add: closedin_contains_derived_set derived_set_of_restrict derived_set_of_subset_topspace) lemma derived_set_subset: "S \ topspace X \ (X derived_set_of S \ S \ closedin X S)" by (simp add: closedin_contains_derived_set) lemma closedin_derived_set: "closedin (subtopology X T) S \ S \ topspace X \ S \ T \ (\x. x \ X derived_set_of S \ x \ T \ x \ S)" by (auto simp: closedin_contains_derived_set derived_set_of_subtopology Int_absorb1) lemma closedin_Int_closure_of: "closedin (subtopology X S) T \ S \ X closure_of T = T" by (metis Int_left_absorb closure_of_eq closure_of_subtopology) lemma closure_of_closedin: "closedin X S \ X closure_of S = S" by (simp add: closure_of_eq) lemma closure_of_eq_diff: "X closure_of S = topspace X - \{T. openin X T \ disjnt S T}" by (auto simp: closure_of_def disjnt_iff) lemma closedin_closure_of [simp]: "closedin X (X closure_of S)" unfolding closure_of_eq_diff by blast lemma closure_of_closure_of [simp]: "X closure_of (X closure_of S) = X closure_of S" by (simp add: closure_of_eq) lemma closure_of_hull: assumes "S \ topspace X" shows "X closure_of S = (closedin X) hull S" proof (rule hull_unique [THEN sym]) show "S \ X closure_of S" by (simp add: closure_of_subset assms) next show "closedin X (X closure_of S)" by simp show "\T. \S \ T; closedin X T\ \ X closure_of S \ T" by (metis closure_of_eq closure_of_mono) qed lemma closure_of_minimal: "\S \ T; closedin X T\ \ (X closure_of S) \ T" by (metis closure_of_eq closure_of_mono) lemma closure_of_minimal_eq: "\S \ topspace X; closedin X T\ \ (X closure_of S) \ T \ S \ T" by (meson closure_of_minimal closure_of_subset subset_trans) lemma closure_of_unique: "\S \ T; closedin X T; \T'. \S \ T'; closedin X T'\ \ T \ T'\ \ X closure_of S = T" by (meson closedin_closure_of closedin_subset closure_of_minimal closure_of_subset eq_iff order.trans) lemma closure_of_eq_empty_gen: "X closure_of S = {} \ disjnt (topspace X) S" unfolding disjnt_def closure_of_restrict [where S=S] using closure_of by fastforce lemma closure_of_eq_empty: "S \ topspace X \ X closure_of S = {} \ S = {}" using closure_of_subset by fastforce lemma openin_Int_closure_of_subset: assumes "openin X S" shows "S \ X closure_of T \ X closure_of (S \ T)" proof - have "S \ X derived_set_of T = S \ X derived_set_of (S \ T)" by (meson assms openin_Int_derived_set_of_eq) moreover have "S \ (S \ T) = S \ T" by fastforce ultimately show ?thesis by (metis closure_of_alt inf.cobounded2 inf_left_commute inf_sup_distrib1) qed lemma closure_of_openin_Int_closure_of: assumes "openin X S" shows "X closure_of (S \ X closure_of T) = X closure_of (S \ T)" proof show "X closure_of (S \ X closure_of T) \ X closure_of (S \ T)" by (simp add: assms closure_of_minimal openin_Int_closure_of_subset) next show "X closure_of (S \ T) \ X closure_of (S \ X closure_of T)" by (metis Int_lower1 Int_subset_iff assms closedin_closure_of closure_of_minimal_eq closure_of_mono inf_le2 le_infI1 openin_subset) qed lemma openin_Int_closure_of_eq: assumes "openin X S" shows "S \ X closure_of T = S \ X closure_of (S \ T)" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (simp add: assms openin_Int_closure_of_subset) show "?rhs \ ?lhs" by (metis closure_of_mono inf_commute inf_le1 inf_mono order_refl) qed lemma openin_Int_closure_of_eq_empty: assumes "openin X S" shows "S \ X closure_of T = {} \ S \ T = {}" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" unfolding disjoint_iff by (meson assms in_closure_of in_mono openin_subset) show "?rhs \ ?lhs" by (simp add: assms openin_Int_closure_of_eq) qed lemma closure_of_openin_Int_superset: "openin X S \ S \ X closure_of T \ X closure_of (S \ T) = X closure_of S" by (metis closure_of_openin_Int_closure_of inf.orderE) lemma closure_of_openin_subtopology_Int_closure_of: assumes S: "openin (subtopology X U) S" and "T \ U" shows "X closure_of (S \ X closure_of T) = X closure_of (S \ T)" (is "?lhs = ?rhs") proof obtain S0 where S0: "openin X S0" "S = S0 \ U" using assms by (auto simp: openin_subtopology) show "?lhs \ ?rhs" proof - have "S0 \ X closure_of T = S0 \ X closure_of (S0 \ T)" by (meson S0(1) openin_Int_closure_of_eq) moreover have "S0 \ T = S0 \ U \ T" using \T \ U\ by fastforce ultimately have "S \ X closure_of T \ X closure_of (S \ T)" using S0(2) by auto then show ?thesis by (meson closedin_closure_of closure_of_minimal) qed next show "?rhs \ ?lhs" proof - have "T \ S \ T \ X derived_set_of T" by force then show ?thesis by (metis Int_subset_iff S closure_of closure_of_mono inf.cobounded2 inf.coboundedI2 inf_commute openin_closedin_eq topspace_subtopology) qed qed lemma closure_of_subtopology_open: "openin X U \ S \ U \ (subtopology X U) closure_of S = U \ X closure_of S" by (metis closure_of_subtopology inf_absorb2 openin_Int_closure_of_eq) lemma discrete_topology_closure_of: "(discrete_topology U) closure_of S = U \ S" by (metis closedin_discrete_topology closure_of_restrict closure_of_unique discrete_topology_unique inf_sup_ord(1) order_refl) text\ Interior with respect to a topological space. \ definition interior_of :: "'a topology \ 'a set \ 'a set" (infixr "interior'_of" 80) where "X interior_of S \ {x. \T. openin X T \ x \ T \ T \ S}" lemma interior_of_restrict: "X interior_of S = X interior_of (topspace X \ S)" using openin_subset by (auto simp: interior_of_def) lemma interior_of_eq: "(X interior_of S = S) \ openin X S" unfolding interior_of_def using openin_subopen by blast lemma interior_of_openin: "openin X S \ X interior_of S = S" by (simp add: interior_of_eq) lemma interior_of_empty [simp]: "X interior_of {} = {}" by (simp add: interior_of_eq) lemma interior_of_topspace [simp]: "X interior_of (topspace X) = topspace X" by (simp add: interior_of_eq) lemma openin_interior_of [simp]: "openin X (X interior_of S)" unfolding interior_of_def using openin_subopen by fastforce lemma interior_of_interior_of [simp]: "X interior_of X interior_of S = X interior_of S" by (simp add: interior_of_eq) lemma interior_of_subset: "X interior_of S \ S" by (auto simp: interior_of_def) lemma interior_of_subset_closure_of: "X interior_of S \ X closure_of S" by (metis closure_of_subset_Int dual_order.trans interior_of_restrict interior_of_subset) lemma subset_interior_of_eq: "S \ X interior_of S \ openin X S" by (metis interior_of_eq interior_of_subset subset_antisym) lemma interior_of_mono: "S \ T \ X interior_of S \ X interior_of T" by (auto simp: interior_of_def) lemma interior_of_maximal: "\T \ S; openin X T\ \ T \ X interior_of S" by (auto simp: interior_of_def) lemma interior_of_maximal_eq: "openin X T \ T \ X interior_of S \ T \ S" by (meson interior_of_maximal interior_of_subset order_trans) lemma interior_of_unique: "\T \ S; openin X T; \T'. \T' \ S; openin X T'\ \ T' \ T\ \ X interior_of S = T" by (simp add: interior_of_maximal_eq interior_of_subset subset_antisym) lemma interior_of_subset_topspace: "X interior_of S \ topspace X" by (simp add: openin_subset) lemma interior_of_subset_subtopology: "(subtopology X S) interior_of T \ S" by (meson openin_imp_subset openin_interior_of) lemma interior_of_Int: "X interior_of (S \ T) = X interior_of S \ X interior_of T" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (simp add: interior_of_mono) show "?rhs \ ?lhs" by (meson inf_mono interior_of_maximal interior_of_subset openin_Int openin_interior_of) qed lemma interior_of_Inter_subset: "X interior_of (\\) \ (\S \ \. X interior_of S)" by (simp add: INT_greatest Inf_lower interior_of_mono) lemma union_interior_of_subset: "X interior_of S \ X interior_of T \ X interior_of (S \ T)" by (simp add: interior_of_mono) lemma interior_of_eq_empty: "X interior_of S = {} \ (\T. openin X T \ T \ S \ T = {})" by (metis bot.extremum_uniqueI interior_of_maximal interior_of_subset openin_interior_of) lemma interior_of_eq_empty_alt: "X interior_of S = {} \ (\T. openin X T \ T \ {} \ T - S \ {})" by (auto simp: interior_of_eq_empty) lemma interior_of_Union_openin_subsets: "\{T. openin X T \ T \ S} = X interior_of S" by (rule interior_of_unique [symmetric]) auto lemma interior_of_complement: "X interior_of (topspace X - S) = topspace X - X closure_of S" by (auto simp: interior_of_def closure_of_def) lemma interior_of_closure_of: "X interior_of S = topspace X - X closure_of (topspace X - S)" unfolding interior_of_complement [symmetric] by (metis Diff_Diff_Int interior_of_restrict) lemma closure_of_interior_of: "X closure_of S = topspace X - X interior_of (topspace X - S)" by (simp add: interior_of_complement Diff_Diff_Int closure_of) lemma closure_of_complement: "X closure_of (topspace X - S) = topspace X - X interior_of S" unfolding interior_of_def closure_of_def by (blast dest: openin_subset) lemma interior_of_eq_empty_complement: "X interior_of S = {} \ X closure_of (topspace X - S) = topspace X" using interior_of_subset_topspace [of X S] closure_of_complement by fastforce lemma closure_of_eq_topspace: "X closure_of S = topspace X \ X interior_of (topspace X - S) = {}" using closure_of_subset_topspace [of X S] interior_of_complement by fastforce lemma interior_of_subtopology_subset: "U \ X interior_of S \ (subtopology X U) interior_of S" by (auto simp: interior_of_def openin_subtopology) lemma interior_of_subtopology_subsets: "T \ U \ T \ (subtopology X U) interior_of S \ (subtopology X T) interior_of S" by (metis inf.absorb_iff2 interior_of_subtopology_subset subtopology_subtopology) lemma interior_of_subtopology_mono: "\S \ T; T \ U\ \ (subtopology X U) interior_of S \ (subtopology X T) interior_of S" by (metis dual_order.trans inf.orderE inf_commute interior_of_subset interior_of_subtopology_subsets) lemma interior_of_subtopology_open: assumes "openin X U" shows "(subtopology X U) interior_of S = U \ X interior_of S" proof - have "\A. U \ X closure_of (U \ A) = U \ X closure_of A" using assms openin_Int_closure_of_eq by blast then have "topspace X \ U - U \ X closure_of (topspace X \ U - S) = U \ (topspace X - X closure_of (topspace X - S))" by (metis (no_types) Diff_Int_distrib Int_Diff inf_commute) then show ?thesis unfolding interior_of_closure_of closure_of_subtopology_open topspace_subtopology using openin_Int_closure_of_eq [OF assms] by (metis assms closure_of_subtopology_open) qed lemma dense_intersects_open: "X closure_of S = topspace X \ (\T. openin X T \ T \ {} \ S \ T \ {})" proof - have "X closure_of S = topspace X \ (topspace X - X interior_of (topspace X - S) = topspace X)" by (simp add: closure_of_interior_of) also have "\ \ X interior_of (topspace X - S) = {}" by (simp add: closure_of_complement interior_of_eq_empty_complement) also have "\ \ (\T. openin X T \ T \ {} \ S \ T \ {})" unfolding interior_of_eq_empty_alt using openin_subset by fastforce finally show ?thesis . qed lemma interior_of_closedin_union_empty_interior_of: assumes "closedin X S" and disj: "X interior_of T = {}" shows "X interior_of (S \ T) = X interior_of S" proof - have "X closure_of (topspace X - T) = topspace X" by (metis Diff_Diff_Int disj closure_of_eq_topspace closure_of_restrict interior_of_closure_of) then show ?thesis unfolding interior_of_closure_of by (metis Diff_Un Diff_subset assms(1) closedin_def closure_of_openin_Int_superset) qed lemma interior_of_union_eq_empty: "closedin X S \ (X interior_of (S \ T) = {} \ X interior_of S = {} \ X interior_of T = {})" by (metis interior_of_closedin_union_empty_interior_of le_sup_iff subset_empty union_interior_of_subset) lemma discrete_topology_interior_of [simp]: "(discrete_topology U) interior_of S = U \ S" by (simp add: interior_of_restrict [of _ S] interior_of_eq) subsection \Frontier with respect to topological space \ definition frontier_of :: "'a topology \ 'a set \ 'a set" (infixr "frontier'_of" 80) where "X frontier_of S \ X closure_of S - X interior_of S" lemma frontier_of_closures: "X frontier_of S = X closure_of S \ X closure_of (topspace X - S)" by (metis Diff_Diff_Int closure_of_complement closure_of_subset_topspace double_diff frontier_of_def interior_of_subset_closure_of) lemma interior_of_union_frontier_of [simp]: "X interior_of S \ X frontier_of S = X closure_of S" by (simp add: frontier_of_def interior_of_subset_closure_of subset_antisym) lemma frontier_of_restrict: "X frontier_of S = X frontier_of (topspace X \ S)" by (metis closure_of_restrict frontier_of_def interior_of_restrict) lemma closedin_frontier_of: "closedin X (X frontier_of S)" by (simp add: closedin_Int frontier_of_closures) lemma frontier_of_subset_topspace: "X frontier_of S \ topspace X" by (simp add: closedin_frontier_of closedin_subset) lemma frontier_of_subset_subtopology: "(subtopology X S) frontier_of T \ S" by (metis (no_types) closedin_derived_set closedin_frontier_of) lemma frontier_of_subtopology_subset: "U \ (subtopology X U) frontier_of S \ (X frontier_of S)" proof - have "U \ X interior_of S - subtopology X U interior_of S = {}" by (simp add: interior_of_subtopology_subset) moreover have "X closure_of S \ subtopology X U closure_of S = subtopology X U closure_of S" by (meson closure_of_subtopology_subset inf.absorb_iff2) ultimately show ?thesis unfolding frontier_of_def by blast qed lemma frontier_of_subtopology_mono: "\S \ T; T \ U\ \ (subtopology X T) frontier_of S \ (subtopology X U) frontier_of S" by (simp add: frontier_of_def Diff_mono closure_of_subtopology_mono interior_of_subtopology_mono) lemma clopenin_eq_frontier_of: "closedin X S \ openin X S \ S \ topspace X \ X frontier_of S = {}" proof (cases "S \ topspace X") case True then show ?thesis by (metis Diff_eq_empty_iff closure_of_eq closure_of_subset_eq frontier_of_def interior_of_eq interior_of_subset interior_of_union_frontier_of sup_bot_right) next case False then show ?thesis by (simp add: frontier_of_closures openin_closedin_eq) qed lemma frontier_of_eq_empty: "S \ topspace X \ (X frontier_of S = {} \ closedin X S \ openin X S)" by (simp add: clopenin_eq_frontier_of) lemma frontier_of_openin: "openin X S \ X frontier_of S = X closure_of S - S" by (metis (no_types) frontier_of_def interior_of_eq) lemma frontier_of_openin_straddle_Int: assumes "openin X U" "U \ X frontier_of S \ {}" shows "U \ S \ {}" "U - S \ {}" proof - have "U \ (X closure_of S \ X closure_of (topspace X - S)) \ {}" using assms by (simp add: frontier_of_closures) then show "U \ S \ {}" using assms openin_Int_closure_of_eq_empty by fastforce show "U - S \ {}" proof - have "\A. X closure_of (A - S) \ U \ {}" using \U \ (X closure_of S \ X closure_of (topspace X - S)) \ {}\ by blast then have "\ U \ S" by (metis Diff_disjoint Diff_eq_empty_iff Int_Diff assms(1) inf_commute openin_Int_closure_of_eq_empty) then show ?thesis by blast qed qed lemma frontier_of_subset_closedin: "closedin X S \ (X frontier_of S) \ S" using closure_of_eq frontier_of_def by fastforce lemma frontier_of_empty [simp]: "X frontier_of {} = {}" by (simp add: frontier_of_def) lemma frontier_of_topspace [simp]: "X frontier_of topspace X = {}" by (simp add: frontier_of_def) lemma frontier_of_subset_eq: assumes "S \ topspace X" shows "(X frontier_of S) \ S \ closedin X S" proof show "X frontier_of S \ S \ closedin X S" by (metis assms closure_of_subset_eq interior_of_subset interior_of_union_frontier_of le_sup_iff) show "closedin X S \ X frontier_of S \ S" by (simp add: frontier_of_subset_closedin) qed lemma frontier_of_complement: "X frontier_of (topspace X - S) = X frontier_of S" by (metis Diff_Diff_Int closure_of_restrict frontier_of_closures inf_commute) lemma frontier_of_disjoint_eq: assumes "S \ topspace X" shows "((X frontier_of S) \ S = {} \ openin X S)" proof assume "X frontier_of S \ S = {}" then have "closedin X (topspace X - S)" using assms closure_of_subset frontier_of_def interior_of_eq interior_of_subset by fastforce then show "openin X S" using assms by (simp add: openin_closedin) next show "openin X S \ X frontier_of S \ S = {}" by (simp add: Diff_Diff_Int closedin_def frontier_of_openin inf.absorb_iff2 inf_commute) qed lemma frontier_of_disjoint_eq_alt: "S \ (topspace X - X frontier_of S) \ openin X S" proof (cases "S \ topspace X") case True show ?thesis using True frontier_of_disjoint_eq by auto next case False then show ?thesis by (meson Diff_subset openin_subset subset_trans) qed lemma frontier_of_Int: "X frontier_of (S \ T) = X closure_of (S \ T) \ (X frontier_of S \ X frontier_of T)" proof - have *: "U \ S \ U \ T \ U \ (S \ A \ T \ B) = U \ (A \ B)" for U S T A B :: "'a set" by blast show ?thesis by (simp add: frontier_of_closures closure_of_mono Diff_Int * flip: closure_of_Un) qed lemma frontier_of_Int_subset: "X frontier_of (S \ T) \ X frontier_of S \ X frontier_of T" by (simp add: frontier_of_Int) lemma frontier_of_Int_closedin: assumes "closedin X S" "closedin X T" shows "X frontier_of(S \ T) = X frontier_of S \ T \ S \ X frontier_of T" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" using assms by (force simp add: frontier_of_Int closedin_Int closure_of_closedin) show "?rhs \ ?lhs" using assms frontier_of_subset_closedin by (auto simp add: frontier_of_Int closedin_Int closure_of_closedin) qed lemma frontier_of_Un_subset: "X frontier_of(S \ T) \ X frontier_of S \ X frontier_of T" by (metis Diff_Un frontier_of_Int_subset frontier_of_complement) lemma frontier_of_Union_subset: "finite \ \ X frontier_of (\\) \ (\T \ \. X frontier_of T)" proof (induction \ rule: finite_induct) case (insert A \) then show ?case using frontier_of_Un_subset by fastforce qed simp lemma frontier_of_frontier_of_subset: "X frontier_of (X frontier_of S) \ X frontier_of S" by (simp add: closedin_frontier_of frontier_of_subset_closedin) lemma frontier_of_subtopology_open: "openin X U \ (subtopology X U) frontier_of S = U \ X frontier_of S" by (simp add: Diff_Int_distrib closure_of_subtopology_open frontier_of_def interior_of_subtopology_open) lemma discrete_topology_frontier_of [simp]: "(discrete_topology U) frontier_of S = {}" by (simp add: Diff_eq discrete_topology_closure_of frontier_of_closures) subsection\Locally finite collections\ definition locally_finite_in where "locally_finite_in X \ \ (\\ \ topspace X) \ (\x \ topspace X. \V. openin X V \ x \ V \ finite {U \ \. U \ V \ {}})" lemma finite_imp_locally_finite_in: "\finite \; \\ \ topspace X\ \ locally_finite_in X \" by (auto simp: locally_finite_in_def) lemma locally_finite_in_subset: assumes "locally_finite_in X \" "\ \ \" shows "locally_finite_in X \" proof - have "finite (\ \ {U. U \ V \ {}}) \ finite (\ \ {U. U \ V \ {}})" for V by (meson \\ \ \\ finite_subset inf_le1 inf_le2 le_inf_iff subset_trans) then show ?thesis using assms unfolding locally_finite_in_def Int_def by fastforce qed lemma locally_finite_in_refinement: assumes \: "locally_finite_in X \" and f: "\S. S \ \ \ f S \ S" shows "locally_finite_in X (f ` \)" proof - show ?thesis unfolding locally_finite_in_def proof safe fix x assume "x \ topspace X" then obtain V where "openin X V" "x \ V" "finite {U \ \. U \ V \ {}}" using \ unfolding locally_finite_in_def by blast moreover have "{U \ \. f U \ V \ {}} \ {U \ \. U \ V \ {}}" for V using f by blast ultimately have "finite {U \ \. f U \ V \ {}}" using finite_subset by blast moreover have "f ` {U \ \. f U \ V \ {}} = {U \ f ` \. U \ V \ {}}" by blast ultimately have "finite {U \ f ` \. U \ V \ {}}" by (metis (no_types, lifting) finite_imageI) then show "\V. openin X V \ x \ V \ finite {U \ f ` \. U \ V \ {}}" using \openin X V\ \x \ V\ by blast next show "\x xa. \xa \ \; x \ f xa\ \ x \ topspace X" by (meson Sup_upper \ f locally_finite_in_def subset_iff) qed qed lemma locally_finite_in_subtopology: assumes \: "locally_finite_in X \" "\\ \ S" shows "locally_finite_in (subtopology X S) \" unfolding locally_finite_in_def proof safe fix x assume x: "x \ topspace (subtopology X S)" then obtain V where "openin X V" "x \ V" and fin: "finite {U \ \. U \ V \ {}}" using \ unfolding locally_finite_in_def topspace_subtopology by blast show "\V. openin (subtopology X S) V \ x \ V \ finite {U \ \. U \ V \ {}}" proof (intro exI conjI) show "openin (subtopology X S) (S \ V)" by (simp add: \openin X V\ openin_subtopology_Int2) have "{U \ \. U \ (S \ V) \ {}} \ {U \ \. U \ V \ {}}" by auto with fin show "finite {U \ \. U \ (S \ V) \ {}}" using finite_subset by auto show "x \ S \ V" using x \x \ V\ by (simp) qed next show "\x A. \x \ A; A \ \\ \ x \ topspace (subtopology X S)" using assms unfolding locally_finite_in_def topspace_subtopology by blast qed lemma closedin_locally_finite_Union: assumes clo: "\S. S \ \ \ closedin X S" and \: "locally_finite_in X \" shows "closedin X (\\)" using \ unfolding locally_finite_in_def closedin_def proof clarify show "openin X (topspace X - \\)" proof (subst openin_subopen, clarify) fix x assume "x \ topspace X" and "x \ \\" then obtain V where "openin X V" "x \ V" and fin: "finite {U \ \. U \ V \ {}}" using \ unfolding locally_finite_in_def by blast let ?T = "V - \{S \ \. S \ V \ {}}" show "\T. openin X T \ x \ T \ T \ topspace X - \\" proof (intro exI conjI) show "openin X ?T" by (metis (no_types, lifting) fin \openin X V\ clo closedin_Union mem_Collect_eq openin_diff) show "x \ ?T" using \x \ \\\ \x \ V\ by auto show "?T \ topspace X - \\" using \openin X V\ openin_subset by auto qed qed qed lemma locally_finite_in_closure: assumes \: "locally_finite_in X \" shows "locally_finite_in X ((\S. X closure_of S) ` \)" using \ unfolding locally_finite_in_def proof (intro conjI; clarsimp) fix x A assume "x \ X closure_of A" then show "x \ topspace X" by (meson in_closure_of) next fix x assume "x \ topspace X" and "\\ \ topspace X" then obtain V where V: "openin X V" "x \ V" and fin: "finite {U \ \. U \ V \ {}}" using \ unfolding locally_finite_in_def by blast have eq: "{y \ f ` \. Q y} = f ` {x. x \ \ \ Q(f x)}" for f Q by blast have eq2: "{A \ \. X closure_of A \ V \ {}} = {A \ \. A \ V \ {}}" using openin_Int_closure_of_eq_empty V by blast have "finite {U \ (closure_of) X ` \. U \ V \ {}}" by (simp add: eq eq2 fin) with V show "\V. openin X V \ x \ V \ finite {U \ (closure_of) X ` \. U \ V \ {}}" by blast qed lemma closedin_Union_locally_finite_closure: "locally_finite_in X \ \ closedin X (\((\S. X closure_of S) ` \))" by (metis (mono_tags) closedin_closure_of closedin_locally_finite_Union imageE locally_finite_in_closure) lemma closure_of_Union_subset: "\((\S. X closure_of S) ` \) \ X closure_of (\\)" by clarify (meson Union_upper closure_of_mono subsetD) lemma closure_of_locally_finite_Union: assumes "locally_finite_in X \" shows "X closure_of (\\) = \((\S. X closure_of S) ` \)" proof (rule closure_of_unique) show "\ \ \ \ ((closure_of) X ` \)" using assms by (simp add: SUP_upper2 Sup_le_iff closure_of_subset locally_finite_in_def) show "closedin X (\ ((closure_of) X ` \))" using assms by (simp add: closedin_Union_locally_finite_closure) show "\T'. \\ \ \ T'; closedin X T'\ \ \ ((closure_of) X ` \) \ T'" by (simp add: Sup_le_iff closure_of_minimal) qed subsection\<^marker>\tag important\ \Continuous maps\ text \We will need to deal with continuous maps in terms of topologies and not in terms of type classes, as defined below.\ definition continuous_map where "continuous_map X Y f \ (\x \ topspace X. f x \ topspace Y) \ (\U. openin Y U \ openin X {x \ topspace X. f x \ U})" lemma continuous_map: "continuous_map X Y f \ f ` (topspace X) \ topspace Y \ (\U. openin Y U \ openin X {x \ topspace X. f x \ U})" by (auto simp: continuous_map_def) lemma continuous_map_image_subset_topspace: "continuous_map X Y f \ f ` (topspace X) \ topspace Y" by (auto simp: continuous_map_def) lemma continuous_map_on_empty: "topspace X = {} \ continuous_map X Y f" by (auto simp: continuous_map_def) lemma continuous_map_closedin: "continuous_map X Y f \ (\x \ topspace X. f x \ topspace Y) \ (\C. closedin Y C \ closedin X {x \ topspace X. f x \ C})" proof - have "(\U. openin Y U \ openin X {x \ topspace X. f x \ U}) = (\C. closedin Y C \ closedin X {x \ topspace X. f x \ C})" if "\x. x \ topspace X \ f x \ topspace Y" proof - have eq: "{x \ topspace X. f x \ topspace Y \ f x \ C} = (topspace X - {x \ topspace X. f x \ C})" for C using that by blast show ?thesis proof (intro iffI allI impI) fix C assume "\U. openin Y U \ openin X {x \ topspace X. f x \ U}" and "closedin Y C" then have "openin X {x \ topspace X. f x \ topspace Y - C}" by blast then show "closedin X {x \ topspace X. f x \ C}" by (auto simp add: closedin_def eq) next fix U assume "\C. closedin Y C \ closedin X {x \ topspace X. f x \ C}" and "openin Y U" then have "closedin X {x \ topspace X. f x \ topspace Y - U}" by blast then show "openin X {x \ topspace X. f x \ U}" by (auto simp add: openin_closedin_eq eq) qed qed then show ?thesis by (auto simp: continuous_map_def) qed lemma openin_continuous_map_preimage: "\continuous_map X Y f; openin Y U\ \ openin X {x \ topspace X. f x \ U}" by (simp add: continuous_map_def) lemma closedin_continuous_map_preimage: "\continuous_map X Y f; closedin Y C\ \ closedin X {x \ topspace X. f x \ C}" by (simp add: continuous_map_closedin) lemma openin_continuous_map_preimage_gen: assumes "continuous_map X Y f" "openin X U" "openin Y V" shows "openin X {x \ U. f x \ V}" proof - have eq: "{x \ U. f x \ V} = U \ {x \ topspace X. f x \ V}" using assms(2) openin_closedin_eq by fastforce show ?thesis unfolding eq using assms openin_continuous_map_preimage by fastforce qed lemma closedin_continuous_map_preimage_gen: assumes "continuous_map X Y f" "closedin X U" "closedin Y V" shows "closedin X {x \ U. f x \ V}" proof - have eq: "{x \ U. f x \ V} = U \ {x \ topspace X. f x \ V}" using assms(2) closedin_def by fastforce show ?thesis unfolding eq using assms closedin_continuous_map_preimage by fastforce qed lemma continuous_map_image_closure_subset: assumes "continuous_map X Y f" shows "f ` (X closure_of S) \ Y closure_of f ` S" proof - have *: "f ` (topspace X) \ topspace Y" by (meson assms continuous_map) have "X closure_of T \ {x \ X closure_of T. f x \ Y closure_of (f ` T)}" if "T \ topspace X" for T proof (rule closure_of_minimal) show "T \ {x \ X closure_of T. f x \ Y closure_of f ` T}" using closure_of_subset * that by (fastforce simp: in_closure_of) next show "closedin X {x \ X closure_of T. f x \ Y closure_of f ` T}" using assms closedin_continuous_map_preimage_gen by fastforce qed then have "f ` (X closure_of (topspace X \ S)) \ Y closure_of (f ` (topspace X \ S))" by blast also have "\ \ Y closure_of (topspace Y \ f ` S)" using * by (blast intro!: closure_of_mono) finally have "f ` (X closure_of (topspace X \ S)) \ Y closure_of (topspace Y \ f ` S)" . then show ?thesis by (metis closure_of_restrict) qed lemma continuous_map_subset_aux1: "continuous_map X Y f \ (\S. f ` (X closure_of S) \ Y closure_of f ` S)" using continuous_map_image_closure_subset by blast lemma continuous_map_subset_aux2: assumes "\S. S \ topspace X \ f ` (X closure_of S) \ Y closure_of f ` S" shows "continuous_map X Y f" unfolding continuous_map_closedin proof (intro conjI ballI allI impI) fix x assume "x \ topspace X" then show "f x \ topspace Y" using assms closure_of_subset_topspace by fastforce next fix C assume "closedin Y C" then show "closedin X {x \ topspace X. f x \ C}" proof (clarsimp simp flip: closure_of_subset_eq, intro conjI) fix x assume x: "x \ X closure_of {x \ topspace X. f x \ C}" and "C \ topspace Y" and "Y closure_of C \ C" show "x \ topspace X" by (meson x in_closure_of) have "{a \ topspace X. f a \ C} \ topspace X" by simp moreover have "Y closure_of f ` {a \ topspace X. f a \ C} \ C" by (simp add: \closedin Y C\ closure_of_minimal image_subset_iff) ultimately have "f ` (X closure_of {a \ topspace X. f a \ C}) \ C" using assms by blast then show "f x \ C" using x by auto qed qed lemma continuous_map_eq_image_closure_subset: "continuous_map X Y f \ (\S. f ` (X closure_of S) \ Y closure_of f ` S)" using continuous_map_subset_aux1 continuous_map_subset_aux2 by metis lemma continuous_map_eq_image_closure_subset_alt: "continuous_map X Y f \ (\S. S \ topspace X \ f ` (X closure_of S) \ Y closure_of f ` S)" using continuous_map_subset_aux1 continuous_map_subset_aux2 by metis lemma continuous_map_eq_image_closure_subset_gen: "continuous_map X Y f \ f ` (topspace X) \ topspace Y \ (\S. f ` (X closure_of S) \ Y closure_of f ` S)" using continuous_map_subset_aux1 continuous_map_subset_aux2 continuous_map_image_subset_topspace by metis lemma continuous_map_closure_preimage_subset: "continuous_map X Y f \ X closure_of {x \ topspace X. f x \ T} \ {x \ topspace X. f x \ Y closure_of T}" unfolding continuous_map_closedin by (rule closure_of_minimal) (use in_closure_of in \fastforce+\) lemma continuous_map_frontier_frontier_preimage_subset: assumes "continuous_map X Y f" shows "X frontier_of {x \ topspace X. f x \ T} \ {x \ topspace X. f x \ Y frontier_of T}" proof - have eq: "topspace X - {x \ topspace X. f x \ T} = {x \ topspace X. f x \ topspace Y - T}" using assms unfolding continuous_map_def by blast have "X closure_of {x \ topspace X. f x \ T} \ {x \ topspace X. f x \ Y closure_of T}" by (simp add: assms continuous_map_closure_preimage_subset) moreover have "X closure_of (topspace X - {x \ topspace X. f x \ T}) \ {x \ topspace X. f x \ Y closure_of (topspace Y - T)}" using continuous_map_closure_preimage_subset [OF assms] eq by presburger ultimately show ?thesis by (auto simp: frontier_of_closures) qed lemma topology_finer_continuous_id: assumes "topspace X = topspace Y" shows "(\S. openin X S \ openin Y S) \ continuous_map Y X id" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" unfolding continuous_map_def using assms openin_subopen openin_subset by fastforce show "?rhs \ ?lhs" unfolding continuous_map_def using assms openin_subopen topspace_def by fastforce qed lemma continuous_map_const [simp]: "continuous_map X Y (\x. C) \ topspace X = {} \ C \ topspace Y" proof (cases "topspace X = {}") case False show ?thesis proof (cases "C \ topspace Y") case True with openin_subopen show ?thesis by (auto simp: continuous_map_def) next case False then show ?thesis unfolding continuous_map_def by fastforce qed qed (auto simp: continuous_map_on_empty) declare continuous_map_const [THEN iffD2, continuous_intros] lemma continuous_map_compose [continuous_intros]: assumes f: "continuous_map X X' f" and g: "continuous_map X' X'' g" shows "continuous_map X X'' (g \ f)" unfolding continuous_map_def proof (intro conjI ballI allI impI) fix x assume "x \ topspace X" then show "(g \ f) x \ topspace X''" using assms unfolding continuous_map_def by force next fix U assume "openin X'' U" have eq: "{x \ topspace X. (g \ f) x \ U} = {x \ topspace X. f x \ {y. y \ topspace X' \ g y \ U}}" by auto (meson f continuous_map_def) show "openin X {x \ topspace X. (g \ f) x \ U}" unfolding eq using assms unfolding continuous_map_def using \openin X'' U\ by blast qed lemma continuous_map_eq: assumes "continuous_map X X' f" and "\x. x \ topspace X \ f x = g x" shows "continuous_map X X' g" proof - have eq: "{x \ topspace X. f x \ U} = {x \ topspace X. g x \ U}" for U using assms by auto show ?thesis using assms by (simp add: continuous_map_def eq) qed lemma restrict_continuous_map [simp]: "topspace X \ S \ continuous_map X X' (restrict f S) \ continuous_map X X' f" by (auto simp: elim!: continuous_map_eq) lemma continuous_map_in_subtopology: "continuous_map X (subtopology X' S) f \ continuous_map X X' f \ f ` (topspace X) \ S" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs proof - have "\A. f ` (X closure_of A) \ subtopology X' S closure_of f ` A" by (meson L continuous_map_image_closure_subset) then show ?thesis by (metis (no_types) closure_of_subset_subtopology closure_of_subtopology_subset closure_of_topspace continuous_map_eq_image_closure_subset dual_order.trans) qed next assume R: ?rhs then have eq: "{x \ topspace X. f x \ U} = {x \ topspace X. f x \ U \ f x \ S}" for U by auto show ?lhs using R unfolding continuous_map by (auto simp: openin_subtopology eq) qed lemma continuous_map_from_subtopology: "continuous_map X X' f \ continuous_map (subtopology X S) X' f" by (auto simp: continuous_map openin_subtopology) lemma continuous_map_into_fulltopology: "continuous_map X (subtopology X' T) f \ continuous_map X X' f" by (auto simp: continuous_map_in_subtopology) lemma continuous_map_into_subtopology: "\continuous_map X X' f; f ` topspace X \ T\ \ continuous_map X (subtopology X' T) f" by (auto simp: continuous_map_in_subtopology) lemma continuous_map_from_subtopology_mono: "\continuous_map (subtopology X T) X' f; S \ T\ \ continuous_map (subtopology X S) X' f" by (metis inf.absorb_iff2 continuous_map_from_subtopology subtopology_subtopology) lemma continuous_map_from_discrete_topology [simp]: "continuous_map (discrete_topology U) X f \ f ` U \ topspace X" by (auto simp: continuous_map_def) lemma continuous_map_iff_continuous [simp]: "continuous_map (top_of_set S) euclidean g = continuous_on S g" by (fastforce simp add: continuous_map openin_subtopology continuous_on_open_invariant) lemma continuous_map_iff_continuous2 [simp]: "continuous_map euclidean euclidean g = continuous_on UNIV g" by (metis continuous_map_iff_continuous subtopology_UNIV) lemma continuous_map_openin_preimage_eq: "continuous_map X Y f \ f ` (topspace X) \ topspace Y \ (\U. openin Y U \ openin X (topspace X \ f -` U))" by (auto simp: continuous_map_def vimage_def Int_def) lemma continuous_map_closedin_preimage_eq: "continuous_map X Y f \ f ` (topspace X) \ topspace Y \ (\U. closedin Y U \ closedin X (topspace X \ f -` U))" by (auto simp: continuous_map_closedin vimage_def Int_def) lemma continuous_map_square_root: "continuous_map euclideanreal euclideanreal sqrt" by (simp add: continuous_at_imp_continuous_on isCont_real_sqrt) lemma continuous_map_sqrt [continuous_intros]: "continuous_map X euclideanreal f \ continuous_map X euclideanreal (\x. sqrt(f x))" by (meson continuous_map_compose continuous_map_eq continuous_map_square_root o_apply) lemma continuous_map_id [simp, continuous_intros]: "continuous_map X X id" unfolding continuous_map_def using openin_subopen topspace_def by fastforce declare continuous_map_id [unfolded id_def, simp, continuous_intros] lemma continuous_map_id_subt [simp]: "continuous_map (subtopology X S) X id" by (simp add: continuous_map_from_subtopology) declare continuous_map_id_subt [unfolded id_def, simp] lemma\<^marker>\tag important\ continuous_map_alt: "continuous_map T1 T2 f = ((\U. openin T2 U \ openin T1 (f -` U \ topspace T1)) \ f ` topspace T1 \ topspace T2)" by (auto simp: continuous_map_def vimage_def image_def Collect_conj_eq inf_commute) lemma continuous_map_open [intro]: "continuous_map T1 T2 f \ openin T2 U \ openin T1 (f-`U \ topspace(T1))" unfolding continuous_map_alt by auto lemma continuous_map_preimage_topspace [intro]: assumes "continuous_map T1 T2 f" shows "f-`(topspace T2) \ topspace T1 = topspace T1" using assms unfolding continuous_map_def by auto subsection\Open and closed maps (not a priori assumed continuous)\ definition open_map :: "'a topology \ 'b topology \ ('a \ 'b) \ bool" where "open_map X1 X2 f \ \U. openin X1 U \ openin X2 (f ` U)" definition closed_map :: "'a topology \ 'b topology \ ('a \ 'b) \ bool" where "closed_map X1 X2 f \ \U. closedin X1 U \ closedin X2 (f ` U)" lemma open_map_imp_subset_topspace: "open_map X1 X2 f \ f ` (topspace X1) \ topspace X2" unfolding open_map_def by (simp add: openin_subset) lemma open_map_on_empty: "topspace X = {} \ open_map X Y f" by (metis empty_iff imageE in_mono open_map_def openin_subopen openin_subset) lemma closed_map_on_empty: "topspace X = {} \ closed_map X Y f" by (simp add: closed_map_def closedin_topspace_empty) lemma closed_map_const: "closed_map X Y (\x. c) \ topspace X = {} \ closedin Y {c}" proof (cases "topspace X = {}") case True then show ?thesis by (simp add: closed_map_on_empty) next case False then show ?thesis by (auto simp: closed_map_def image_constant_conv) qed lemma open_map_imp_subset: "\open_map X1 X2 f; S \ topspace X1\ \ f ` S \ topspace X2" by (meson order_trans open_map_imp_subset_topspace subset_image_iff) lemma topology_finer_open_id: "(\S. openin X S \ openin X' S) \ open_map X X' id" unfolding open_map_def by auto lemma open_map_id: "open_map X X id" unfolding open_map_def by auto lemma open_map_eq: "\open_map X X' f; \x. x \ topspace X \ f x = g x\ \ open_map X X' g" unfolding open_map_def by (metis image_cong openin_subset subset_iff) lemma open_map_inclusion_eq: "open_map (subtopology X S) X id \ openin X (topspace X \ S)" proof - have *: "openin X (T \ S)" if "openin X (S \ topspace X)" "openin X T" for T proof - have "T \ topspace X" using that by (simp add: openin_subset) with that show "openin X (T \ S)" by (metis inf.absorb1 inf.left_commute inf_commute openin_Int) qed show ?thesis by (fastforce simp add: open_map_def Int_commute openin_subtopology_alt intro: *) qed lemma open_map_inclusion: "openin X S \ open_map (subtopology X S) X id" by (simp add: open_map_inclusion_eq openin_Int) lemma open_map_compose: "\open_map X X' f; open_map X' X'' g\ \ open_map X X'' (g \ f)" by (metis (no_types, lifting) image_comp open_map_def) lemma closed_map_imp_subset_topspace: "closed_map X1 X2 f \ f ` (topspace X1) \ topspace X2" by (simp add: closed_map_def closedin_subset) lemma closed_map_imp_subset: "\closed_map X1 X2 f; S \ topspace X1\ \ f ` S \ topspace X2" using closed_map_imp_subset_topspace by blast lemma topology_finer_closed_id: "(\S. closedin X S \ closedin X' S) \ closed_map X X' id" by (simp add: closed_map_def) lemma closed_map_id: "closed_map X X id" by (simp add: closed_map_def) lemma closed_map_eq: "\closed_map X X' f; \x. x \ topspace X \ f x = g x\ \ closed_map X X' g" unfolding closed_map_def by (metis image_cong closedin_subset subset_iff) lemma closed_map_compose: "\closed_map X X' f; closed_map X' X'' g\ \ closed_map X X'' (g \ f)" by (metis (no_types, lifting) closed_map_def image_comp) lemma closed_map_inclusion_eq: "closed_map (subtopology X S) X id \ closedin X (topspace X \ S)" proof - have *: "closedin X (T \ S)" if "closedin X (S \ topspace X)" "closedin X T" for T proof - have "T \ topspace X" using that by (simp add: closedin_subset) with that show "closedin X (T \ S)" by (metis inf.absorb1 inf.left_commute inf_commute closedin_Int) qed show ?thesis by (fastforce simp add: closed_map_def Int_commute closedin_subtopology_alt intro: *) qed lemma closed_map_inclusion: "closedin X S \ closed_map (subtopology X S) X id" by (simp add: closed_map_inclusion_eq closedin_Int) lemma open_map_into_subtopology: "\open_map X X' f; f ` topspace X \ S\ \ open_map X (subtopology X' S) f" unfolding open_map_def openin_subtopology using openin_subset by fastforce lemma closed_map_into_subtopology: "\closed_map X X' f; f ` topspace X \ S\ \ closed_map X (subtopology X' S) f" unfolding closed_map_def closedin_subtopology using closedin_subset by fastforce lemma open_map_into_discrete_topology: "open_map X (discrete_topology U) f \ f ` (topspace X) \ U" unfolding open_map_def openin_discrete_topology using openin_subset by blast lemma closed_map_into_discrete_topology: "closed_map X (discrete_topology U) f \ f ` (topspace X) \ U" unfolding closed_map_def closedin_discrete_topology using closedin_subset by blast lemma bijective_open_imp_closed_map: "\open_map X X' f; f ` (topspace X) = topspace X'; inj_on f (topspace X)\ \ closed_map X X' f" unfolding open_map_def closed_map_def closedin_def by auto (metis Diff_subset inj_on_image_set_diff) lemma bijective_closed_imp_open_map: "\closed_map X X' f; f ` (topspace X) = topspace X'; inj_on f (topspace X)\ \ open_map X X' f" unfolding closed_map_def open_map_def openin_closedin_eq by auto (metis Diff_subset inj_on_image_set_diff) lemma open_map_from_subtopology: "\open_map X X' f; openin X U\ \ open_map (subtopology X U) X' f" unfolding open_map_def openin_subtopology_alt by blast lemma closed_map_from_subtopology: "\closed_map X X' f; closedin X U\ \ closed_map (subtopology X U) X' f" unfolding closed_map_def closedin_subtopology_alt by blast lemma open_map_restriction: assumes f: "open_map X X' f" and U: "{x \ topspace X. f x \ V} = U" shows "open_map (subtopology X U) (subtopology X' V) f" unfolding open_map_def proof clarsimp fix W assume "openin (subtopology X U) W" then obtain T where "openin X T" "W = T \ U" by (meson openin_subtopology) with f U have "f ` W = (f ` T) \ V" unfolding open_map_def openin_closedin_eq by auto then show "openin (subtopology X' V) (f ` W)" by (metis \openin X T\ f open_map_def openin_subtopology_Int) qed lemma closed_map_restriction: assumes f: "closed_map X X' f" and U: "{x \ topspace X. f x \ V} = U" shows "closed_map (subtopology X U) (subtopology X' V) f" unfolding closed_map_def proof clarsimp fix W assume "closedin (subtopology X U) W" then obtain T where "closedin X T" "W = T \ U" by (meson closedin_subtopology) with f U have "f ` W = (f ` T) \ V" unfolding closed_map_def closedin_def by auto then show "closedin (subtopology X' V) (f ` W)" by (metis \closedin X T\ closed_map_def closedin_subtopology f) qed subsection\Quotient maps\ definition quotient_map where "quotient_map X X' f \ f ` (topspace X) = topspace X' \ (\U. U \ topspace X' \ (openin X {x. x \ topspace X \ f x \ U} \ openin X' U))" lemma quotient_map_eq: assumes "quotient_map X X' f" "\x. x \ topspace X \ f x = g x" shows "quotient_map X X' g" proof - have eq: "{x \ topspace X. f x \ U} = {x \ topspace X. g x \ U}" for U using assms by auto show ?thesis using assms unfolding quotient_map_def by (metis (mono_tags, lifting) eq image_cong) qed lemma quotient_map_compose: assumes f: "quotient_map X X' f" and g: "quotient_map X' X'' g" shows "quotient_map X X'' (g \ f)" unfolding quotient_map_def proof (intro conjI allI impI) show "(g \ f) ` topspace X = topspace X''" using assms by (simp only: image_comp [symmetric]) (simp add: quotient_map_def) next fix U'' assume "U'' \ topspace X''" define U' where "U' \ {y \ topspace X'. g y \ U''}" have "U' \ topspace X'" by (auto simp add: U'_def) then have U': "openin X {x \ topspace X. f x \ U'} = openin X' U'" using assms unfolding quotient_map_def by simp have eq: "{x \ topspace X. f x \ topspace X' \ g (f x) \ U''} = {x \ topspace X. (g \ f) x \ U''}" using f quotient_map_def by fastforce have "openin X {x \ topspace X. (g \ f) x \ U''} = openin X {x \ topspace X. f x \ U'}" using assms by (simp add: quotient_map_def U'_def eq) also have "\ = openin X'' U''" using U'_def \U'' \ topspace X''\ U' g quotient_map_def by fastforce finally show "openin X {x \ topspace X. (g \ f) x \ U''} = openin X'' U''" . qed lemma quotient_map_from_composition: assumes f: "continuous_map X X' f" and g: "continuous_map X' X'' g" and gf: "quotient_map X X'' (g \ f)" shows "quotient_map X' X'' g" unfolding quotient_map_def proof (intro conjI allI impI) show "g ` topspace X' = topspace X''" using assms unfolding continuous_map_def quotient_map_def by fastforce next fix U'' :: "'c set" assume U'': "U'' \ topspace X''" have eq: "{x \ topspace X. g (f x) \ U''} = {x \ topspace X. f x \ {y. y \ topspace X' \ g y \ U''}}" using continuous_map_def f by fastforce show "openin X' {x \ topspace X'. g x \ U''} = openin X'' U''" using assms unfolding continuous_map_def quotient_map_def by (metis (mono_tags, lifting) Collect_cong U'' comp_apply eq) qed lemma quotient_imp_continuous_map: "quotient_map X X' f \ continuous_map X X' f" by (simp add: continuous_map openin_subset quotient_map_def) lemma quotient_imp_surjective_map: "quotient_map X X' f \ f ` (topspace X) = topspace X'" by (simp add: quotient_map_def) lemma quotient_map_closedin: "quotient_map X X' f \ f ` (topspace X) = topspace X' \ (\U. U \ topspace X' \ (closedin X {x. x \ topspace X \ f x \ U} \ closedin X' U))" proof - have eq: "(topspace X - {x \ topspace X. f x \ U'}) = {x \ topspace X. f x \ topspace X' \ f x \ U'}" if "f ` topspace X = topspace X'" "U' \ topspace X'" for U' using that by auto have "(\U\topspace X'. openin X {x \ topspace X. f x \ U} = openin X' U) = (\U\topspace X'. closedin X {x \ topspace X. f x \ U} = closedin X' U)" if "f ` topspace X = topspace X'" proof (rule iffI; intro allI impI subsetI) fix U' assume *[rule_format]: "\U\topspace X'. openin X {x \ topspace X. f x \ U} = openin X' U" and U': "U' \ topspace X'" show "closedin X {x \ topspace X. f x \ U'} = closedin X' U'" using U' by (auto simp add: closedin_def simp flip: * [of "topspace X' - U'"] eq [OF that]) next fix U' :: "'b set" assume *[rule_format]: "\U\topspace X'. closedin X {x \ topspace X. f x \ U} = closedin X' U" and U': "U' \ topspace X'" show "openin X {x \ topspace X. f x \ U'} = openin X' U'" using U' by (auto simp add: openin_closedin_eq simp flip: * [of "topspace X' - U'"] eq [OF that]) qed then show ?thesis unfolding quotient_map_def by force qed lemma continuous_open_imp_quotient_map: assumes "continuous_map X X' f" and om: "open_map X X' f" and feq: "f ` (topspace X) = topspace X'" shows "quotient_map X X' f" proof - { fix U assume U: "U \ topspace X'" and "openin X {x \ topspace X. f x \ U}" then have ope: "openin X' (f ` {x \ topspace X. f x \ U})" using om unfolding open_map_def by blast then have "openin X' U" using U feq by (subst openin_subopen) force } moreover have "openin X {x \ topspace X. f x \ U}" if "U \ topspace X'" and "openin X' U" for U using that assms unfolding continuous_map_def by blast ultimately show ?thesis unfolding quotient_map_def using assms by blast qed lemma continuous_closed_imp_quotient_map: assumes "continuous_map X X' f" and om: "closed_map X X' f" and feq: "f ` (topspace X) = topspace X'" shows "quotient_map X X' f" proof - have "f ` {x \ topspace X. f x \ U} = U" if "U \ topspace X'" for U using that feq by auto with assms show ?thesis unfolding quotient_map_closedin closed_map_def continuous_map_closedin by auto qed lemma continuous_open_quotient_map: "\continuous_map X X' f; open_map X X' f\ \ quotient_map X X' f \ f ` (topspace X) = topspace X'" by (meson continuous_open_imp_quotient_map quotient_map_def) lemma continuous_closed_quotient_map: "\continuous_map X X' f; closed_map X X' f\ \ quotient_map X X' f \ f ` (topspace X) = topspace X'" by (meson continuous_closed_imp_quotient_map quotient_map_def) lemma injective_quotient_map: assumes "inj_on f (topspace X)" shows "quotient_map X X' f \ continuous_map X X' f \ open_map X X' f \ closed_map X X' f \ f ` (topspace X) = topspace X'" (is "?lhs = ?rhs") proof assume L: ?lhs have "open_map X X' f" proof (clarsimp simp add: open_map_def) fix U assume "openin X U" then have "U \ topspace X" by (simp add: openin_subset) moreover have "{x \ topspace X. f x \ f ` U} = U" using \U \ topspace X\ assms inj_onD by fastforce ultimately show "openin X' (f ` U)" using L unfolding quotient_map_def by (metis (no_types, lifting) Collect_cong \openin X U\ image_mono) qed moreover have "closed_map X X' f" proof (clarsimp simp add: closed_map_def) fix U assume "closedin X U" then have "U \ topspace X" by (simp add: closedin_subset) moreover have "{x \ topspace X. f x \ f ` U} = U" using \U \ topspace X\ assms inj_onD by fastforce ultimately show "closedin X' (f ` U)" using L unfolding quotient_map_closedin by (metis (no_types, lifting) Collect_cong \closedin X U\ image_mono) qed ultimately show ?rhs using L by (simp add: quotient_imp_continuous_map quotient_imp_surjective_map) next assume ?rhs then show ?lhs by (simp add: continuous_closed_imp_quotient_map) qed lemma continuous_compose_quotient_map: assumes f: "quotient_map X X' f" and g: "continuous_map X X'' (g \ f)" shows "continuous_map X' X'' g" unfolding quotient_map_def continuous_map_def proof (intro conjI ballI allI impI) show "\x'. x' \ topspace X' \ g x' \ topspace X''" using assms unfolding quotient_map_def by (metis (no_types, opaque_lifting) continuous_map_image_subset_topspace image_comp image_subset_iff) next fix U'' :: "'c set" assume U'': "openin X'' U''" have "f ` topspace X = topspace X'" by (simp add: f quotient_imp_surjective_map) then have eq: "{x \ topspace X. f x \ topspace X' \ g (f x) \ U} = {x \ topspace X. g (f x) \ U}" for U by auto have "openin X {x \ topspace X. f x \ topspace X' \ g (f x) \ U''}" unfolding eq using U'' g openin_continuous_map_preimage by fastforce then have *: "openin X {x \ topspace X. f x \ {x \ topspace X'. g x \ U''}}" by auto show "openin X' {x \ topspace X'. g x \ U''}" using f unfolding quotient_map_def by (metis (no_types) Collect_subset *) qed lemma continuous_compose_quotient_map_eq: "quotient_map X X' f \ continuous_map X X'' (g \ f) \ continuous_map X' X'' g" using continuous_compose_quotient_map continuous_map_compose quotient_imp_continuous_map by blast lemma quotient_map_compose_eq: "quotient_map X X' f \ quotient_map X X'' (g \ f) \ quotient_map X' X'' g" by (meson continuous_compose_quotient_map_eq quotient_imp_continuous_map quotient_map_compose quotient_map_from_composition) lemma quotient_map_restriction: assumes quo: "quotient_map X Y f" and U: "{x \ topspace X. f x \ V} = U" and disj: "openin Y V \ closedin Y V" shows "quotient_map (subtopology X U) (subtopology Y V) f" using disj proof assume V: "openin Y V" with U have sub: "U \ topspace X" "V \ topspace Y" by (auto simp: openin_subset) have fim: "f ` topspace X = topspace Y" and Y: "\U. U \ topspace Y \ openin X {x \ topspace X. f x \ U} = openin Y U" using quo unfolding quotient_map_def by auto have "openin X U" using U V Y sub(2) by blast show ?thesis unfolding quotient_map_def proof (intro conjI allI impI) show "f ` topspace (subtopology X U) = topspace (subtopology Y V)" using sub U fim by (auto) next fix Y' :: "'b set" assume "Y' \ topspace (subtopology Y V)" then have "Y' \ topspace Y" "Y' \ V" by (simp_all) then have eq: "{x \ topspace X. x \ U \ f x \ Y'} = {x \ topspace X. f x \ Y'}" using U by blast then show "openin (subtopology X U) {x \ topspace (subtopology X U). f x \ Y'} = openin (subtopology Y V) Y'" using U V Y \openin X U\ \Y' \ topspace Y\ \Y' \ V\ by (simp add: openin_open_subtopology eq) (auto simp: openin_closedin_eq) qed next assume V: "closedin Y V" with U have sub: "U \ topspace X" "V \ topspace Y" by (auto simp: closedin_subset) have fim: "f ` topspace X = topspace Y" and Y: "\U. U \ topspace Y \ closedin X {x \ topspace X. f x \ U} = closedin Y U" using quo unfolding quotient_map_closedin by auto have "closedin X U" using U V Y sub(2) by blast show ?thesis unfolding quotient_map_closedin proof (intro conjI allI impI) show "f ` topspace (subtopology X U) = topspace (subtopology Y V)" using sub U fim by (auto) next fix Y' :: "'b set" assume "Y' \ topspace (subtopology Y V)" then have "Y' \ topspace Y" "Y' \ V" by (simp_all) then have eq: "{x \ topspace X. x \ U \ f x \ Y'} = {x \ topspace X. f x \ Y'}" using U by blast then show "closedin (subtopology X U) {x \ topspace (subtopology X U). f x \ Y'} = closedin (subtopology Y V) Y'" using U V Y \closedin X U\ \Y' \ topspace Y\ \Y' \ V\ by (simp add: closedin_closed_subtopology eq) (auto simp: closedin_def) qed qed lemma quotient_map_saturated_open: "quotient_map X Y f \ continuous_map X Y f \ f ` (topspace X) = topspace Y \ (\U. openin X U \ {x \ topspace X. f x \ f ` U} \ U \ openin Y (f ` U))" (is "?lhs = ?rhs") proof assume L: ?lhs then have fim: "f ` topspace X = topspace Y" and Y: "\U. U \ topspace Y \ openin Y U = openin X {x \ topspace X. f x \ U}" unfolding quotient_map_def by auto show ?rhs proof (intro conjI allI impI) show "continuous_map X Y f" by (simp add: L quotient_imp_continuous_map) show "f ` topspace X = topspace Y" by (simp add: fim) next fix U :: "'a set" assume U: "openin X U \ {x \ topspace X. f x \ f ` U} \ U" then have sub: "f ` U \ topspace Y" and eq: "{x \ topspace X. f x \ f ` U} = U" using fim openin_subset by fastforce+ show "openin Y (f ` U)" by (simp add: sub Y eq U) qed next assume ?rhs then have YX: "\U. openin Y U \ openin X {x \ topspace X. f x \ U}" and fim: "f ` topspace X = topspace Y" and XY: "\U. \openin X U; {x \ topspace X. f x \ f ` U} \ U\ \ openin Y (f ` U)" by (auto simp: quotient_map_def continuous_map_def) show ?lhs proof (simp add: quotient_map_def fim, intro allI impI iffI) fix U :: "'b set" assume "U \ topspace Y" and X: "openin X {x \ topspace X. f x \ U}" have feq: "f ` {x \ topspace X. f x \ U} = U" using \U \ topspace Y\ fim by auto show "openin Y U" using XY [OF X] by (simp add: feq) next fix U :: "'b set" assume "U \ topspace Y" and Y: "openin Y U" show "openin X {x \ topspace X. f x \ U}" by (metis YX [OF Y]) qed qed subsection\ Separated Sets\ definition separatedin :: "'a topology \ 'a set \ 'a set \ bool" where "separatedin X S T \ S \ topspace X \ T \ topspace X \ S \ X closure_of T = {} \ T \ X closure_of S = {}" lemma separatedin_empty [simp]: "separatedin X S {} \ S \ topspace X" "separatedin X {} S \ S \ topspace X" by (simp_all add: separatedin_def) lemma separatedin_refl [simp]: "separatedin X S S \ S = {}" proof - have "\x. \separatedin X S S; x \ S\ \ False" by (metis all_not_in_conv closure_of_subset inf.orderE separatedin_def) then show ?thesis by auto qed lemma separatedin_sym: "separatedin X S T \ separatedin X T S" by (auto simp: separatedin_def) lemma separatedin_imp_disjoint: "separatedin X S T \ disjnt S T" by (meson closure_of_subset disjnt_def disjnt_subset2 separatedin_def) lemma separatedin_mono: "\separatedin X S T; S' \ S; T' \ T\ \ separatedin X S' T'" unfolding separatedin_def using closure_of_mono by blast lemma separatedin_open_sets: "\openin X S; openin X T\ \ separatedin X S T \ disjnt S T" unfolding disjnt_def separatedin_def by (auto simp: openin_Int_closure_of_eq_empty openin_subset) lemma separatedin_closed_sets: "\closedin X S; closedin X T\ \ separatedin X S T \ disjnt S T" unfolding closure_of_eq disjnt_def separatedin_def by (metis closedin_def closure_of_eq inf_commute) lemma separatedin_subtopology: "separatedin (subtopology X U) S T \ S \ U \ T \ U \ separatedin X S T" (is "?lhs = ?rhs") by (auto simp: separatedin_def closure_of_subtopology Int_ac disjoint_iff elim!: inf.orderE) lemma separatedin_discrete_topology: "separatedin (discrete_topology U) S T \ S \ U \ T \ U \ disjnt S T" by (metis openin_discrete_topology separatedin_def separatedin_open_sets topspace_discrete_topology) lemma separated_eq_distinguishable: "separatedin X {x} {y} \ x \ topspace X \ y \ topspace X \ (\U. openin X U \ x \ U \ (y \ U)) \ (\v. openin X v \ y \ v \ (x \ v))" by (force simp: separatedin_def closure_of_def) lemma separatedin_Un [simp]: "separatedin X S (T \ U) \ separatedin X S T \ separatedin X S U" "separatedin X (S \ T) U \ separatedin X S U \ separatedin X T U" by (auto simp: separatedin_def) lemma separatedin_Union: "finite \ \ separatedin X S (\\) \ S \ topspace X \ (\T \ \. separatedin X S T)" "finite \ \ separatedin X (\\) S \ (\T \ \. separatedin X S T) \ S \ topspace X" by (auto simp: separatedin_def closure_of_Union) lemma separatedin_openin_diff: "\openin X S; openin X T\ \ separatedin X (S - T) (T - S)" unfolding separatedin_def by (metis Diff_Int_distrib2 Diff_disjoint Diff_empty Diff_mono empty_Diff empty_subsetI openin_Int_closure_of_eq_empty openin_subset) lemma separatedin_closedin_diff: assumes "closedin X S" "closedin X T" shows "separatedin X (S - T) (T - S)" proof - have "S - T \ topspace X" "T - S \ topspace X" using assms closedin_subset by auto with assms show ?thesis by (simp add: separatedin_def Diff_Int_distrib2 closure_of_minimal inf_absorb2) qed lemma separation_closedin_Un_gen: "separatedin X S T \ S \ topspace X \ T \ topspace X \ disjnt S T \ closedin (subtopology X (S \ T)) S \ closedin (subtopology X (S \ T)) T" by (auto simp add: separatedin_def closedin_Int_closure_of disjnt_iff dest: closure_of_subset) lemma separation_openin_Un_gen: "separatedin X S T \ S \ topspace X \ T \ topspace X \ disjnt S T \ openin (subtopology X (S \ T)) S \ openin (subtopology X (S \ T)) T" unfolding openin_closedin_eq topspace_subtopology separation_closedin_Un_gen disjnt_def by (auto simp: Diff_triv Int_commute Un_Diff inf_absorb1 topspace_def) subsection\Homeomorphisms\ text\(1-way and 2-way versions may be useful in places)\ definition homeomorphic_map :: "'a topology \ 'b topology \ ('a \ 'b) \ bool" where "homeomorphic_map X Y f \ quotient_map X Y f \ inj_on f (topspace X)" definition homeomorphic_maps :: "'a topology \ 'b topology \ ('a \ 'b) \ ('b \ 'a) \ bool" where "homeomorphic_maps X Y f g \ continuous_map X Y f \ continuous_map Y X g \ (\x \ topspace X. g(f x) = x) \ (\y \ topspace Y. f(g y) = y)" lemma homeomorphic_map_eq: "\homeomorphic_map X Y f; \x. x \ topspace X \ f x = g x\ \ homeomorphic_map X Y g" by (meson homeomorphic_map_def inj_on_cong quotient_map_eq) lemma homeomorphic_maps_eq: "\homeomorphic_maps X Y f g; \x. x \ topspace X \ f x = f' x; \y. y \ topspace Y \ g y = g' y\ \ homeomorphic_maps X Y f' g'" unfolding homeomorphic_maps_def by (metis continuous_map_eq continuous_map_eq_image_closure_subset_gen image_subset_iff) lemma homeomorphic_maps_sym: "homeomorphic_maps X Y f g \ homeomorphic_maps Y X g f" by (auto simp: homeomorphic_maps_def) lemma homeomorphic_maps_id: "homeomorphic_maps X Y id id \ Y = X" (is "?lhs = ?rhs") proof assume L: ?lhs then have "topspace X = topspace Y" by (auto simp: homeomorphic_maps_def continuous_map_def) with L show ?rhs unfolding homeomorphic_maps_def by (metis topology_finer_continuous_id topology_eq) next assume ?rhs then show ?lhs unfolding homeomorphic_maps_def by auto qed lemma homeomorphic_map_id [simp]: "homeomorphic_map X Y id \ Y = X" (is "?lhs = ?rhs") proof assume L: ?lhs then have eq: "topspace X = topspace Y" by (auto simp: homeomorphic_map_def continuous_map_def quotient_map_def) then have "\S. openin X S \ openin Y S" by (meson L homeomorphic_map_def injective_quotient_map topology_finer_open_id) then show ?rhs using L unfolding homeomorphic_map_def by (metis eq quotient_imp_continuous_map topology_eq topology_finer_continuous_id) next assume ?rhs then show ?lhs unfolding homeomorphic_map_def by (simp add: closed_map_id continuous_closed_imp_quotient_map) qed lemma homeomorphic_map_compose: assumes "homeomorphic_map X Y f" "homeomorphic_map Y X'' g" shows "homeomorphic_map X X'' (g \ f)" proof - have "inj_on g (f ` topspace X)" by (metis (no_types) assms homeomorphic_map_def quotient_imp_surjective_map) then show ?thesis using assms by (meson comp_inj_on homeomorphic_map_def quotient_map_compose_eq) qed lemma homeomorphic_maps_compose: "homeomorphic_maps X Y f h \ homeomorphic_maps Y X'' g k \ homeomorphic_maps X X'' (g \ f) (h \ k)" unfolding homeomorphic_maps_def by (auto simp: continuous_map_compose; simp add: continuous_map_def) lemma homeomorphic_eq_everything_map: "homeomorphic_map X Y f \ continuous_map X Y f \ open_map X Y f \ closed_map X Y f \ f ` (topspace X) = topspace Y \ inj_on f (topspace X)" unfolding homeomorphic_map_def by (force simp: injective_quotient_map intro: injective_quotient_map) lemma homeomorphic_imp_continuous_map: "homeomorphic_map X Y f \ continuous_map X Y f" by (simp add: homeomorphic_eq_everything_map) lemma homeomorphic_imp_open_map: "homeomorphic_map X Y f \ open_map X Y f" by (simp add: homeomorphic_eq_everything_map) lemma homeomorphic_imp_closed_map: "homeomorphic_map X Y f \ closed_map X Y f" by (simp add: homeomorphic_eq_everything_map) lemma homeomorphic_imp_surjective_map: "homeomorphic_map X Y f \ f ` (topspace X) = topspace Y" by (simp add: homeomorphic_eq_everything_map) lemma homeomorphic_imp_injective_map: "homeomorphic_map X Y f \ inj_on f (topspace X)" by (simp add: homeomorphic_eq_everything_map) lemma bijective_open_imp_homeomorphic_map: "\continuous_map X Y f; open_map X Y f; f ` (topspace X) = topspace Y; inj_on f (topspace X)\ \ homeomorphic_map X Y f" by (simp add: homeomorphic_map_def continuous_open_imp_quotient_map) lemma bijective_closed_imp_homeomorphic_map: "\continuous_map X Y f; closed_map X Y f; f ` (topspace X) = topspace Y; inj_on f (topspace X)\ \ homeomorphic_map X Y f" by (simp add: continuous_closed_quotient_map homeomorphic_map_def) lemma open_eq_continuous_inverse_map: assumes X: "\x. x \ topspace X \ f x \ topspace Y \ g(f x) = x" and Y: "\y. y \ topspace Y \ g y \ topspace X \ f(g y) = y" shows "open_map X Y f \ continuous_map Y X g" proof - have eq: "{x \ topspace Y. g x \ U} = f ` U" if "openin X U" for U using openin_subset [OF that] by (force simp: X Y image_iff) show ?thesis by (auto simp: Y open_map_def continuous_map_def eq) qed lemma closed_eq_continuous_inverse_map: assumes X: "\x. x \ topspace X \ f x \ topspace Y \ g(f x) = x" and Y: "\y. y \ topspace Y \ g y \ topspace X \ f(g y) = y" shows "closed_map X Y f \ continuous_map Y X g" proof - have eq: "{x \ topspace Y. g x \ U} = f ` U" if "closedin X U" for U using closedin_subset [OF that] by (force simp: X Y image_iff) show ?thesis by (auto simp: Y closed_map_def continuous_map_closedin eq) qed lemma homeomorphic_maps_map: "homeomorphic_maps X Y f g \ homeomorphic_map X Y f \ homeomorphic_map Y X g \ (\x \ topspace X. g(f x) = x) \ (\y \ topspace Y. f(g y) = y)" (is "?lhs = ?rhs") proof assume ?lhs then have L: "continuous_map X Y f" "continuous_map Y X g" "\x\topspace X. g (f x) = x" "\x'\topspace Y. f (g x') = x'" by (auto simp: homeomorphic_maps_def) show ?rhs proof (intro conjI bijective_open_imp_homeomorphic_map L) show "open_map X Y f" using L using open_eq_continuous_inverse_map [of concl: X Y f g] by (simp add: continuous_map_def) show "open_map Y X g" using L using open_eq_continuous_inverse_map [of concl: Y X g f] by (simp add: continuous_map_def) show "f ` topspace X = topspace Y" "g ` topspace Y = topspace X" using L by (force simp: continuous_map_closedin)+ show "inj_on f (topspace X)" "inj_on g (topspace Y)" using L unfolding inj_on_def by metis+ qed next assume ?rhs then show ?lhs by (auto simp: homeomorphic_maps_def homeomorphic_imp_continuous_map) qed lemma homeomorphic_maps_imp_map: "homeomorphic_maps X Y f g \ homeomorphic_map X Y f" using homeomorphic_maps_map by blast lemma homeomorphic_map_maps: "homeomorphic_map X Y f \ (\g. homeomorphic_maps X Y f g)" (is "?lhs = ?rhs") proof assume ?lhs then have L: "continuous_map X Y f" "open_map X Y f" "closed_map X Y f" "f ` (topspace X) = topspace Y" "inj_on f (topspace X)" by (auto simp: homeomorphic_eq_everything_map) have X: "\x. x \ topspace X \ f x \ topspace Y \ inv_into (topspace X) f (f x) = x" using L by auto have Y: "\y. y \ topspace Y \ inv_into (topspace X) f y \ topspace X \ f (inv_into (topspace X) f y) = y" by (simp add: L f_inv_into_f inv_into_into) have "homeomorphic_maps X Y f (inv_into (topspace X) f)" unfolding homeomorphic_maps_def proof (intro conjI L) show "continuous_map Y X (inv_into (topspace X) f)" by (simp add: L X Y flip: open_eq_continuous_inverse_map [where f=f]) next show "\x\topspace X. inv_into (topspace X) f (f x) = x" "\y\topspace Y. f (inv_into (topspace X) f y) = y" using X Y by auto qed then show ?rhs by metis next assume ?rhs then show ?lhs using homeomorphic_maps_map by blast qed lemma homeomorphic_maps_involution: "\continuous_map X X f; \x. x \ topspace X \ f(f x) = x\ \ homeomorphic_maps X X f f" by (auto simp: homeomorphic_maps_def) lemma homeomorphic_map_involution: "\continuous_map X X f; \x. x \ topspace X \ f(f x) = x\ \ homeomorphic_map X X f" using homeomorphic_maps_involution homeomorphic_maps_map by blast lemma homeomorphic_map_openness: assumes hom: "homeomorphic_map X Y f" and U: "U \ topspace X" shows "openin Y (f ` U) \ openin X U" proof - obtain g where "homeomorphic_maps X Y f g" using assms by (auto simp: homeomorphic_map_maps) then have g: "homeomorphic_map Y X g" and gf: "\x. x \ topspace X \ g(f x) = x" by (auto simp: homeomorphic_maps_map) then have "openin X U \ openin Y (f ` U)" using hom homeomorphic_imp_open_map open_map_def by blast show "openin Y (f ` U) = openin X U" proof assume L: "openin Y (f ` U)" have "U = g ` (f ` U)" using U gf by force then show "openin X U" by (metis L homeomorphic_imp_open_map open_map_def g) next assume "openin X U" then show "openin Y (f ` U)" using hom homeomorphic_imp_open_map open_map_def by blast qed qed lemma homeomorphic_map_closedness: assumes hom: "homeomorphic_map X Y f" and U: "U \ topspace X" shows "closedin Y (f ` U) \ closedin X U" proof - obtain g where "homeomorphic_maps X Y f g" using assms by (auto simp: homeomorphic_map_maps) then have g: "homeomorphic_map Y X g" and gf: "\x. x \ topspace X \ g(f x) = x" by (auto simp: homeomorphic_maps_map) then have "closedin X U \ closedin Y (f ` U)" using hom homeomorphic_imp_closed_map closed_map_def by blast show "closedin Y (f ` U) = closedin X U" proof assume L: "closedin Y (f ` U)" have "U = g ` (f ` U)" using U gf by force then show "closedin X U" by (metis L homeomorphic_imp_closed_map closed_map_def g) next assume "closedin X U" then show "closedin Y (f ` U)" using hom homeomorphic_imp_closed_map closed_map_def by blast qed qed lemma homeomorphic_map_openness_eq: "homeomorphic_map X Y f \ openin X U \ U \ topspace X \ openin Y (f ` U)" by (meson homeomorphic_map_openness openin_closedin_eq) lemma homeomorphic_map_closedness_eq: "homeomorphic_map X Y f \ closedin X U \ U \ topspace X \ closedin Y (f ` U)" by (meson closedin_subset homeomorphic_map_closedness) lemma all_openin_homeomorphic_image: assumes "homeomorphic_map X Y f" shows "(\V. openin Y V \ P V) \ (\U. openin X U \ P(f ` U))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (meson assms homeomorphic_map_openness_eq) next assume ?rhs then show ?lhs by (metis (no_types, lifting) assms homeomorphic_imp_surjective_map homeomorphic_map_openness openin_subset subset_image_iff) qed lemma all_closedin_homeomorphic_image: assumes "homeomorphic_map X Y f" shows "(\V. closedin Y V \ P V) \ (\U. closedin X U \ P(f ` U))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (meson assms homeomorphic_map_closedness_eq) next assume ?rhs then show ?lhs by (metis (no_types, lifting) assms homeomorphic_imp_surjective_map homeomorphic_map_closedness closedin_subset subset_image_iff) qed lemma homeomorphic_map_derived_set_of: assumes hom: "homeomorphic_map X Y f" and S: "S \ topspace X" shows "Y derived_set_of (f ` S) = f ` (X derived_set_of S)" proof - have fim: "f ` (topspace X) = topspace Y" and inj: "inj_on f (topspace X)" using hom by (auto simp: homeomorphic_eq_everything_map) have iff: "(\T. x \ T \ openin X T \ (\y. y \ x \ y \ S \ y \ T)) = (\T. T \ topspace Y \ f x \ T \ openin Y T \ (\y. y \ f x \ y \ f ` S \ y \ T))" if "x \ topspace X" for x proof - have \
: "(x \ T \ openin X T) = (T \ topspace X \ f x \ f ` T \ openin Y (f ` T))" for T by (meson hom homeomorphic_map_openness_eq inj inj_on_image_mem_iff that) moreover have "(\y. y \ x \ y \ S \ y \ T) = (\y. y \ f x \ y \ f ` S \ y \ f ` T)" (is "?lhs = ?rhs") if "T \ topspace X \ f x \ f ` T \ openin Y (f ` T)" for T proof show "?lhs \ ?rhs" by (meson \
imageI inj inj_on_eq_iff inj_on_subset that) show "?rhs \ ?lhs" using S inj inj_onD that by fastforce qed ultimately show ?thesis by (auto simp flip: fim simp: all_subset_image) qed have *: "\T = f ` S; \x. x \ S \ P x \ Q(f x)\ \ {y. y \ T \ Q y} = f ` {x \ S. P x}" for T S P Q by auto show ?thesis unfolding derived_set_of_def by (rule *) (use fim iff openin_subset in force)+ qed lemma homeomorphic_map_closure_of: assumes hom: "homeomorphic_map X Y f" and S: "S \ topspace X" shows "Y closure_of (f ` S) = f ` (X closure_of S)" unfolding closure_of using homeomorphic_imp_surjective_map [OF hom] S by (auto simp: in_derived_set_of homeomorphic_map_derived_set_of [OF assms]) lemma homeomorphic_map_interior_of: assumes hom: "homeomorphic_map X Y f" and S: "S \ topspace X" shows "Y interior_of (f ` S) = f ` (X interior_of S)" proof - { fix y assume "y \ topspace Y" and "y \ Y closure_of (topspace Y - f ` S)" then have "y \ f ` (topspace X - X closure_of (topspace X - S))" using homeomorphic_eq_everything_map [THEN iffD1, OF hom] homeomorphic_map_closure_of [OF hom] by (metis DiffI Diff_subset S closure_of_subset_topspace inj_on_image_set_diff) } moreover { fix x assume "x \ topspace X" then have "f x \ topspace Y" using hom homeomorphic_imp_surjective_map by blast } moreover { fix x assume "x \ topspace X" and "x \ X closure_of (topspace X - S)" and "f x \ Y closure_of (topspace Y - f ` S)" then have "False" using homeomorphic_map_closure_of [OF hom] hom unfolding homeomorphic_eq_everything_map by (metis Diff_subset S closure_of_subset_topspace inj_on_image_mem_iff inj_on_image_set_diff) } ultimately show ?thesis by (auto simp: interior_of_closure_of) qed lemma homeomorphic_map_frontier_of: assumes hom: "homeomorphic_map X Y f" and S: "S \ topspace X" shows "Y frontier_of (f ` S) = f ` (X frontier_of S)" unfolding frontier_of_def proof (intro equalityI subsetI DiffI) fix y assume "y \ Y closure_of f ` S - Y interior_of f ` S" then show "y \ f ` (X closure_of S - X interior_of S)" using S hom homeomorphic_map_closure_of homeomorphic_map_interior_of by fastforce next fix y assume "y \ f ` (X closure_of S - X interior_of S)" then show "y \ Y closure_of f ` S" using S hom homeomorphic_map_closure_of by fastforce next fix x assume "x \ f ` (X closure_of S - X interior_of S)" then obtain y where y: "x = f y" "y \ X closure_of S" "y \ X interior_of S" by blast then have "y \ topspace X" by (simp add: in_closure_of) then have "f y \ f ` (X interior_of S)" by (meson hom homeomorphic_map_def inj_on_image_mem_iff interior_of_subset_topspace y(3)) then show "x \ Y interior_of f ` S" using S hom homeomorphic_map_interior_of y(1) by blast qed lemma homeomorphic_maps_subtopologies: "\homeomorphic_maps X Y f g; f ` (topspace X \ S) = topspace Y \ T\ \ homeomorphic_maps (subtopology X S) (subtopology Y T) f g" unfolding homeomorphic_maps_def by (force simp: continuous_map_from_subtopology continuous_map_in_subtopology) lemma homeomorphic_maps_subtopologies_alt: "\homeomorphic_maps X Y f g; f ` (topspace X \ S) \ T; g ` (topspace Y \ T) \ S\ \ homeomorphic_maps (subtopology X S) (subtopology Y T) f g" unfolding homeomorphic_maps_def by (force simp: continuous_map_from_subtopology continuous_map_in_subtopology) lemma homeomorphic_map_subtopologies: "\homeomorphic_map X Y f; f ` (topspace X \ S) = topspace Y \ T\ \ homeomorphic_map (subtopology X S) (subtopology Y T) f" by (meson homeomorphic_map_maps homeomorphic_maps_subtopologies) lemma homeomorphic_map_subtopologies_alt: assumes hom: "homeomorphic_map X Y f" and S: "\x. \x \ topspace X; f x \ topspace Y\ \ f x \ T \ x \ S" shows "homeomorphic_map (subtopology X S) (subtopology Y T) f" proof - have "homeomorphic_maps (subtopology X S) (subtopology Y T) f g" if "homeomorphic_maps X Y f g" for g proof (rule homeomorphic_maps_subtopologies [OF that]) show "f ` (topspace X \ S) = topspace Y \ T" using that S apply (auto simp: homeomorphic_maps_def continuous_map_def) by (metis IntI image_iff) qed then show ?thesis using hom by (meson homeomorphic_map_maps) qed subsection\Relation of homeomorphism between topological spaces\ definition homeomorphic_space (infixr "homeomorphic'_space" 50) where "X homeomorphic_space Y \ \f g. homeomorphic_maps X Y f g" lemma homeomorphic_space_refl: "X homeomorphic_space X" by (meson homeomorphic_maps_id homeomorphic_space_def) lemma homeomorphic_space_sym: "X homeomorphic_space Y \ Y homeomorphic_space X" unfolding homeomorphic_space_def by (metis homeomorphic_maps_sym) lemma homeomorphic_space_trans [trans]: "\X1 homeomorphic_space X2; X2 homeomorphic_space X3\ \ X1 homeomorphic_space X3" unfolding homeomorphic_space_def by (metis homeomorphic_maps_compose) lemma homeomorphic_space: "X homeomorphic_space Y \ (\f. homeomorphic_map X Y f)" by (simp add: homeomorphic_map_maps homeomorphic_space_def) lemma homeomorphic_maps_imp_homeomorphic_space: "homeomorphic_maps X Y f g \ X homeomorphic_space Y" unfolding homeomorphic_space_def by metis lemma homeomorphic_map_imp_homeomorphic_space: "homeomorphic_map X Y f \ X homeomorphic_space Y" unfolding homeomorphic_map_maps using homeomorphic_space_def by blast lemma homeomorphic_empty_space: "X homeomorphic_space Y \ topspace X = {} \ topspace Y = {}" by (metis homeomorphic_imp_surjective_map homeomorphic_space image_is_empty) lemma homeomorphic_empty_space_eq: assumes "topspace X = {}" shows "X homeomorphic_space Y \ topspace Y = {}" proof - have "\f t. continuous_map X (t::'b topology) f" using assms continuous_map_on_empty by blast then show ?thesis by (metis (no_types) assms continuous_map_on_empty empty_iff homeomorphic_empty_space homeomorphic_maps_def homeomorphic_space_def) qed subsection\Connected topological spaces\ definition connected_space :: "'a topology \ bool" where "connected_space X \ \(\E1 E2. openin X E1 \ openin X E2 \ topspace X \ E1 \ E2 \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" definition connectedin :: "'a topology \ 'a set \ bool" where "connectedin X S \ S \ topspace X \ connected_space (subtopology X S)" lemma connected_spaceD: "\connected_space X; openin X U; openin X V; topspace X \ U \ V; U \ V = {}; U \ {}; V \ {}\ \ False" by (auto simp: connected_space_def) lemma connectedin_subset_topspace: "connectedin X S \ S \ topspace X" by (simp add: connectedin_def) lemma connectedin_topspace: "connectedin X (topspace X) \ connected_space X" by (simp add: connectedin_def) lemma connected_space_subtopology: "connectedin X S \ connected_space (subtopology X S)" by (simp add: connectedin_def) lemma connectedin_subtopology: "connectedin (subtopology X S) T \ connectedin X T \ T \ S" by (force simp: connectedin_def subtopology_subtopology inf_absorb2) lemma connected_space_eq: "connected_space X \ (\E1 E2. openin X E1 \ openin X E2 \ E1 \ E2 = topspace X \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" unfolding connected_space_def by (metis openin_Un openin_subset subset_antisym) lemma connected_space_closedin: "connected_space X \ (\E1 E2. closedin X E1 \ closedin X E2 \ topspace X \ E1 \ E2 \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" (is "?lhs = ?rhs") proof assume ?lhs then have L: "\E1 E2. \openin X E1; E1 \ E2 = {}; topspace X \ E1 \ E2; openin X E2\ \ E1 = {} \ E2 = {}" by (simp add: connected_space_def) show ?rhs unfolding connected_space_def proof clarify fix E1 E2 assume "closedin X E1" and "closedin X E2" and "topspace X \ E1 \ E2" and "E1 \ E2 = {}" and "E1 \ {}" and "E2 \ {}" have "E1 \ E2 = topspace X" by (meson Un_subset_iff \closedin X E1\ \closedin X E2\ \topspace X \ E1 \ E2\ closedin_def subset_antisym) then have "topspace X - E2 = E1" using \E1 \ E2 = {}\ by fastforce then have "topspace X = E1" using \E1 \ {}\ L \closedin X E1\ \closedin X E2\ by blast then show "False" using \E1 \ E2 = {}\ \E1 \ E2 = topspace X\ \E2 \ {}\ by blast qed next assume R: ?rhs show ?lhs unfolding connected_space_def proof clarify fix E1 E2 assume "openin X E1" and "openin X E2" and "topspace X \ E1 \ E2" and "E1 \ E2 = {}" and "E1 \ {}" and "E2 \ {}" have "E1 \ E2 = topspace X" by (meson Un_subset_iff \openin X E1\ \openin X E2\ \topspace X \ E1 \ E2\ openin_closedin_eq subset_antisym) then have "topspace X - E2 = E1" using \E1 \ E2 = {}\ by fastforce then have "topspace X = E1" using \E1 \ {}\ R \openin X E1\ \openin X E2\ by blast then show "False" using \E1 \ E2 = {}\ \E1 \ E2 = topspace X\ \E2 \ {}\ by blast qed qed lemma connected_space_closedin_eq: "connected_space X \ (\E1 E2. closedin X E1 \ closedin X E2 \ E1 \ E2 = topspace X \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" by (metis closedin_Un closedin_def connected_space_closedin subset_antisym) lemma connected_space_clopen_in: "connected_space X \ (\T. openin X T \ closedin X T \ T = {} \ T = topspace X)" proof - have eq: "openin X E1 \ openin X E2 \ E1 \ E2 = topspace X \ E1 \ E2 = {} \ P \ E2 = topspace X - E1 \ openin X E1 \ openin X E2 \ P" for E1 E2 P using openin_subset by blast show ?thesis unfolding connected_space_eq eq closedin_def by (auto simp: openin_closedin_eq) qed lemma connectedin: "connectedin X S \ S \ topspace X \ (\E1 E2. openin X E1 \ openin X E2 \ S \ E1 \ E2 \ E1 \ E2 \ S = {} \ E1 \ S \ {} \ E2 \ S \ {})" (is "?lhs = ?rhs") proof - have *: "(\E1:: 'a set. \E2:: 'a set. (\T1:: 'a set. P1 T1 \ E1 = f1 T1) \ (\T2:: 'a set. P2 T2 \ E2 = f2 T2) \ R E1 E2) \ (\T1 T2. P1 T1 \ P2 T2 \ R(f1 T1) (f2 T2))" for P1 f1 P2 f2 R by auto show ?thesis unfolding connectedin_def connected_space_def openin_subtopology topspace_subtopology * by (intro conj_cong arg_cong [where f=Not] ex_cong1; blast dest!: openin_subset) qed lemma connectedin_iff_connected [simp]: "connectedin euclidean S \ connected S" by (simp add: connected_def connectedin) lemma connectedin_closedin: "connectedin X S \ S \ topspace X \ \(\E1 E2. closedin X E1 \ closedin X E2 \ S \ (E1 \ E2) \ (E1 \ E2 \ S = {}) \ \(E1 \ S = {}) \ \(E2 \ S = {}))" proof - have *: "(\E1:: 'a set. \E2:: 'a set. (\T1:: 'a set. P1 T1 \ E1 = f1 T1) \ (\T2:: 'a set. P2 T2 \ E2 = f2 T2) \ R E1 E2) \ (\T1 T2. P1 T1 \ P2 T2 \ R(f1 T1) (f2 T2))" for P1 f1 P2 f2 R by auto show ?thesis unfolding connectedin_def connected_space_closedin closedin_subtopology topspace_subtopology * by (intro conj_cong arg_cong [where f=Not] ex_cong1; blast dest!: openin_subset) qed lemma connectedin_empty [simp]: "connectedin X {}" by (simp add: connectedin) lemma connected_space_topspace_empty: "topspace X = {} \ connected_space X" using connectedin_topspace by fastforce lemma connectedin_sing [simp]: "connectedin X {a} \ a \ topspace X" by (simp add: connectedin) lemma connectedin_absolute [simp]: "connectedin (subtopology X S) S \ connectedin X S" by (simp add: connectedin_subtopology) lemma connectedin_Union: assumes \: "\S. S \ \ \ connectedin X S" and ne: "\\ \ {}" shows "connectedin X (\\)" proof - have "\\ \ topspace X" using \ by (simp add: Union_least connectedin_def) moreover have False if "openin X E1" "openin X E2" and cover: "\\ \ E1 \ E2" and disj: "E1 \ E2 \ \\ = {}" and overlap1: "E1 \ \\ \ {}" and overlap2: "E2 \ \\ \ {}" for E1 E2 proof - have disjS: "E1 \ E2 \ S = {}" if "S \ \" for S using Diff_triv that disj by auto have coverS: "S \ E1 \ E2" if "S \ \" for S using that cover by blast have "\ \ {}" using overlap1 by blast obtain a where a: "\U. U \ \ \ a \ U" using ne by force with \\ \ {}\ have "a \ \\" by blast then consider "a \ E1" | "a \ E2" using \\\ \ E1 \ E2\ by auto then show False proof cases case 1 then obtain b S where "b \ E2" "b \ S" "S \ \" using overlap2 by blast then show ?thesis using "1" \openin X E1\ \openin X E2\ disjS coverS a [OF \S \ \\] \[OF \S \ \\] unfolding connectedin by (meson disjoint_iff_not_equal) next case 2 then obtain b S where "b \ E1" "b \ S" "S \ \" using overlap1 by blast then show ?thesis using "2" \openin X E1\ \openin X E2\ disjS coverS a [OF \S \ \\] \[OF \S \ \\] unfolding connectedin by (meson disjoint_iff_not_equal) qed qed ultimately show ?thesis unfolding connectedin by blast qed lemma connectedin_Un: "\connectedin X S; connectedin X T; S \ T \ {}\ \ connectedin X (S \ T)" using connectedin_Union [of "{S,T}"] by auto lemma connected_space_subconnected: "connected_space X \ (\x \ topspace X. \y \ topspace X. \S. connectedin X S \ x \ S \ y \ S)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using connectedin_topspace by blast next assume R [rule_format]: ?rhs have False if "openin X U" "openin X V" and disj: "U \ V = {}" and cover: "topspace X \ U \ V" and "U \ {}" "V \ {}" for U V proof - obtain u v where "u \ U" "v \ V" using \U \ {}\ \V \ {}\ by auto then obtain T where "u \ T" "v \ T" and T: "connectedin X T" using R [of u v] that by (meson \openin X U\ \openin X V\ subsetD openin_subset) then show False using that unfolding connectedin by (metis IntI \u \ U\ \v \ V\ empty_iff inf_bot_left subset_trans) qed then show ?lhs by (auto simp: connected_space_def) qed lemma connectedin_intermediate_closure_of: assumes "connectedin X S" "S \ T" "T \ X closure_of S" shows "connectedin X T" proof - have S: "S \ topspace X" and T: "T \ topspace X" using assms by (meson closure_of_subset_topspace dual_order.trans)+ have \
: "\E1 E2. \openin X E1; openin X E2; E1 \ S = {} \ E2 \ S = {}\ \ E1 \ T = {} \ E2 \ T = {}" using assms unfolding disjoint_iff by (meson in_closure_of subsetD) then show ?thesis using assms unfolding connectedin closure_of_subset_topspace S T by (metis Int_empty_right T dual_order.trans inf.orderE inf_left_commute) qed lemma connectedin_closure_of: "connectedin X S \ connectedin X (X closure_of S)" by (meson closure_of_subset connectedin_def connectedin_intermediate_closure_of subset_refl) lemma connectedin_separation: "connectedin X S \ S \ topspace X \ (\C1 C2. C1 \ C2 = S \ C1 \ {} \ C2 \ {} \ C1 \ X closure_of C2 = {} \ C2 \ X closure_of C1 = {})" (is "?lhs = ?rhs") unfolding connectedin_def connected_space_closedin_eq closedin_Int_closure_of topspace_subtopology apply (intro conj_cong refl arg_cong [where f=Not]) apply (intro ex_cong1 iffI, blast) using closure_of_subset_Int by force lemma connectedin_eq_not_separated: "connectedin X S \ S \ topspace X \ (\C1 C2. C1 \ C2 = S \ C1 \ {} \ C2 \ {} \ separatedin X C1 C2)" unfolding separatedin_def by (metis connectedin_separation sup.boundedE) lemma connectedin_eq_not_separated_subset: "connectedin X S \ S \ topspace X \ (\C1 C2. S \ C1 \ C2 \ S \ C1 \ {} \ S \ C2 \ {} \ separatedin X C1 C2)" proof - have "\C1 C2. S \ C1 \ C2 \ S \ C1 = {} \ S \ C2 = {} \ \ separatedin X C1 C2" if "\C1 C2. C1 \ C2 = S \ C1 = {} \ C2 = {} \ \ separatedin X C1 C2" proof (intro allI) fix C1 C2 show "S \ C1 \ C2 \ S \ C1 = {} \ S \ C2 = {} \ \ separatedin X C1 C2" using that [of "S \ C1" "S \ C2"] by (auto simp: separatedin_mono) qed then show ?thesis by (metis Un_Int_eq(1) Un_Int_eq(2) connectedin_eq_not_separated order_refl) qed lemma connected_space_eq_not_separated: "connected_space X \ (\C1 C2. C1 \ C2 = topspace X \ C1 \ {} \ C2 \ {} \ separatedin X C1 C2)" by (simp add: connectedin_eq_not_separated flip: connectedin_topspace) lemma connected_space_eq_not_separated_subset: "connected_space X \ (\C1 C2. topspace X \ C1 \ C2 \ C1 \ {} \ C2 \ {} \ separatedin X C1 C2)" by (metis connected_space_eq_not_separated le_sup_iff separatedin_def subset_antisym) lemma connectedin_subset_separated_union: "\connectedin X C; separatedin X S T; C \ S \ T\ \ C \ S \ C \ T" unfolding connectedin_eq_not_separated_subset by blast lemma connectedin_nonseparated_union: assumes "connectedin X S" "connectedin X T" "\separatedin X S T" shows "connectedin X (S \ T)" proof - have "\C1 C2. \T \ C1 \ C2; S \ C1 \ C2\ \ S \ C1 = {} \ T \ C1 = {} \ S \ C2 = {} \ T \ C2 = {} \ \ separatedin X C1 C2" using assms unfolding connectedin_eq_not_separated_subset by (metis (no_types, lifting) assms connectedin_subset_separated_union inf.orderE separatedin_empty(1) separatedin_mono separatedin_sym) then show ?thesis unfolding connectedin_eq_not_separated_subset by (simp add: assms(1) assms(2) connectedin_subset_topspace Int_Un_distrib2) qed lemma connected_space_closures: "connected_space X \ (\e1 e2. e1 \ e2 = topspace X \ X closure_of e1 \ X closure_of e2 = {} \ e1 \ {} \ e2 \ {})" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs unfolding connected_space_closedin_eq by (metis Un_upper1 Un_upper2 closedin_closure_of closure_of_Un closure_of_eq_empty closure_of_topspace) next assume ?rhs then show ?lhs unfolding connected_space_closedin_eq by (metis closure_of_eq) qed lemma connectedin_inter_frontier_of: assumes "connectedin X S" "S \ T \ {}" "S - T \ {}" shows "S \ X frontier_of T \ {}" proof - have "S \ topspace X" and *: "\E1 E2. openin X E1 \ openin X E2 \ E1 \ E2 \ S = {} \ S \ E1 \ E2 \ E1 \ S = {} \ E2 \ S = {}" using \connectedin X S\ by (auto simp: connectedin) moreover have "S - (topspace X \ T) \ {}" using assms(3) by blast moreover have "S \ topspace X \ T \ {}" using assms(1) assms(2) connectedin by fastforce moreover have False if "S \ T \ {}" "S - T \ {}" "T \ topspace X" "S \ X frontier_of T = {}" for T proof - have null: "S \ (X closure_of T - X interior_of T) = {}" using that unfolding frontier_of_def by blast have "X interior_of T \ (topspace X - X closure_of T) \ S = {}" by (metis Diff_disjoint inf_bot_left interior_of_Int interior_of_complement interior_of_empty) moreover have "S \ X interior_of T \ (topspace X - X closure_of T)" using that \S \ topspace X\ null by auto moreover have "S \ X interior_of T \ {}" using closure_of_subset that(1) that(3) null by fastforce ultimately have "S \ X interior_of (topspace X - T) = {}" by (metis "*" inf_commute interior_of_complement openin_interior_of) then have "topspace (subtopology X S) \ X interior_of T = S" using \S \ topspace X\ interior_of_complement null by fastforce then show ?thesis using that by (metis Diff_eq_empty_iff inf_le2 interior_of_subset subset_trans) qed ultimately show ?thesis by (metis Int_lower1 frontier_of_restrict inf_assoc) qed lemma connectedin_continuous_map_image: assumes f: "continuous_map X Y f" and "connectedin X S" shows "connectedin Y (f ` S)" proof - have "S \ topspace X" and *: "\E1 E2. openin X E1 \ openin X E2 \ E1 \ E2 \ S = {} \ S \ E1 \ E2 \ E1 \ S = {} \ E2 \ S = {}" using \connectedin X S\ by (auto simp: connectedin) show ?thesis unfolding connectedin connected_space_def proof (intro conjI notI; clarify) show "f x \ topspace Y" if "x \ S" for x using \S \ topspace X\ continuous_map_image_subset_topspace f that by blast next fix U V let ?U = "{x \ topspace X. f x \ U}" let ?V = "{x \ topspace X. f x \ V}" assume UV: "openin Y U" "openin Y V" "f ` S \ U \ V" "U \ V \ f ` S = {}" "U \ f ` S \ {}" "V \ f ` S \ {}" then have 1: "?U \ ?V \ S = {}" by auto have 2: "openin X ?U" "openin X ?V" using \openin Y U\ \openin Y V\ continuous_map f by fastforce+ show "False" using * [of ?U ?V] UV \S \ topspace X\ by (auto simp: 1 2) qed qed lemma homeomorphic_connected_space: "X homeomorphic_space Y \ connected_space X \ connected_space Y" unfolding homeomorphic_space_def homeomorphic_maps_def by (metis connected_space_subconnected connectedin_continuous_map_image connectedin_topspace continuous_map_image_subset_topspace image_eqI image_subset_iff) lemma homeomorphic_map_connectedness: assumes f: "homeomorphic_map X Y f" and U: "U \ topspace X" shows "connectedin Y (f ` U) \ connectedin X U" proof - have 1: "f ` U \ topspace Y \ U \ topspace X" using U f homeomorphic_imp_surjective_map by blast moreover have "connected_space (subtopology Y (f ` U)) \ connected_space (subtopology X U)" proof (rule homeomorphic_connected_space) have "f ` U \ topspace Y" by (simp add: U 1) then have "topspace Y \ f ` U = f ` U" by (simp add: subset_antisym) then show "subtopology Y (f ` U) homeomorphic_space subtopology X U" by (metis (no_types) Int_subset_iff U f homeomorphic_map_imp_homeomorphic_space homeomorphic_map_subtopologies homeomorphic_space_sym subset_antisym subset_refl) qed ultimately show ?thesis by (auto simp: connectedin_def) qed lemma homeomorphic_map_connectedness_eq: "homeomorphic_map X Y f \ connectedin X U \ U \ topspace X \ connectedin Y (f ` U)" using homeomorphic_map_connectedness connectedin_subset_topspace by metis lemma connectedin_discrete_topology: "connectedin (discrete_topology U) S \ S \ U \ (\a. S \ {a})" proof (cases "S \ U") case True show ?thesis proof (cases "S = {}") case False moreover have "connectedin (discrete_topology U) S \ (\a. S = {a})" proof show "connectedin (discrete_topology U) S \ \a. S = {a}" using False connectedin_inter_frontier_of insert_Diff by fastforce qed (use True in auto) ultimately show ?thesis by auto qed simp next case False then show ?thesis by (simp add: connectedin_def) qed lemma connected_space_discrete_topology: "connected_space (discrete_topology U) \ (\a. U \ {a})" by (metis connectedin_discrete_topology connectedin_topspace order_refl topspace_discrete_topology) subsection\Compact sets\ definition compactin where "compactin X S \ S \ topspace X \ (\\. (\U \ \. openin X U) \ S \ \\ \ (\\. finite \ \ \ \ \ \ S \ \\))" definition compact_space where "compact_space X \ compactin X (topspace X)" lemma compact_space_alt: "compact_space X \ (\\. (\U \ \. openin X U) \ topspace X \ \\ \ (\\. finite \ \ \ \ \ \ topspace X \ \\))" by (simp add: compact_space_def compactin_def) lemma compact_space: "compact_space X \ (\\. (\U \ \. openin X U) \ \\ = topspace X \ (\\. finite \ \ \ \ \ \ \\ = topspace X))" unfolding compact_space_alt using openin_subset by fastforce lemma compactinD: "\compactin X S; \U. U \ \ \ openin X U; S \ \\\ \ \\. finite \ \ \ \ \ \ S \ \\" by (auto simp: compactin_def) lemma compactin_euclidean_iff [simp]: "compactin euclidean S \ compact S" by (simp add: compact_eq_Heine_Borel compactin_def) meson lemma compactin_absolute [simp]: "compactin (subtopology X S) S \ compactin X S" proof - have eq: "(\U \ \. \Y. openin X Y \ U = Y \ S) \ \ \ (\Y. Y \ S) ` {y. openin X y}" for \ by auto show ?thesis by (auto simp: compactin_def openin_subtopology eq imp_conjL all_subset_image ex_finite_subset_image) qed lemma compactin_subspace: "compactin X S \ S \ topspace X \ compact_space (subtopology X S)" unfolding compact_space_def topspace_subtopology by (metis compactin_absolute compactin_def inf.absorb2) lemma compact_space_subtopology: "compactin X S \ compact_space (subtopology X S)" by (simp add: compactin_subspace) lemma compactin_subtopology: "compactin (subtopology X S) T \ compactin X T \ T \ S" by (metis compactin_subspace inf.absorb_iff2 le_inf_iff subtopology_subtopology topspace_subtopology) lemma compactin_subset_topspace: "compactin X S \ S \ topspace X" by (simp add: compactin_subspace) lemma compactin_contractive: "\compactin X' S; topspace X' = topspace X; \U. openin X U \ openin X' U\ \ compactin X S" by (simp add: compactin_def) lemma finite_imp_compactin: "\S \ topspace X; finite S\ \ compactin X S" by (metis compactin_subspace compact_space finite_UnionD inf.absorb_iff2 order_refl topspace_subtopology) lemma compactin_empty [iff]: "compactin X {}" by (simp add: finite_imp_compactin) lemma compact_space_topspace_empty: "topspace X = {} \ compact_space X" by (simp add: compact_space_def) lemma finite_imp_compactin_eq: "finite S \ (compactin X S \ S \ topspace X)" using compactin_subset_topspace finite_imp_compactin by blast lemma compactin_sing [simp]: "compactin X {a} \ a \ topspace X" by (simp add: finite_imp_compactin_eq) lemma closed_compactin: assumes XK: "compactin X K" and "C \ K" and XC: "closedin X C" shows "compactin X C" unfolding compactin_def proof (intro conjI allI impI) show "C \ topspace X" by (simp add: XC closedin_subset) next fix \ :: "'a set set" assume \: "Ball \ (openin X) \ C \ \\" have "(\U\insert (topspace X - C) \. openin X U)" using XC \ by blast moreover have "K \ \(insert (topspace X - C) \)" using \ XK compactin_subset_topspace by fastforce ultimately obtain \ where "finite \" "\ \ insert (topspace X - C) \" "K \ \\" using assms unfolding compactin_def by metis moreover have "openin X (topspace X - C)" using XC by auto ultimately show "\\. finite \ \ \ \ \ \ C \ \\" using \C \ K\ by (rule_tac x="\ - {topspace X - C}" in exI) auto qed lemma closedin_compact_space: "\compact_space X; closedin X S\ \ compactin X S" by (simp add: closed_compactin closedin_subset compact_space_def) lemma compact_Int_closedin: assumes "compactin X S" "closedin X T" shows "compactin X (S \ T)" proof - have "compactin (subtopology X S) (S \ T)" by (metis assms closedin_compact_space closedin_subtopology compactin_subspace inf_commute) then show ?thesis by (simp add: compactin_subtopology) qed lemma closed_Int_compactin: "\closedin X S; compactin X T\ \ compactin X (S \ T)" by (metis compact_Int_closedin inf_commute) lemma compactin_Un: assumes S: "compactin X S" and T: "compactin X T" shows "compactin X (S \ T)" unfolding compactin_def proof (intro conjI allI impI) show "S \ T \ topspace X" using assms by (auto simp: compactin_def) next fix \ :: "'a set set" assume \: "Ball \ (openin X) \ S \ T \ \\" with S obtain \ where \: "finite \" "\ \ \" "S \ \\" unfolding compactin_def by (meson sup.bounded_iff) obtain \ where "finite \" "\ \ \" "T \ \\" using \ T unfolding compactin_def by (meson sup.bounded_iff) with \ show "\\. finite \ \ \ \ \ \ S \ T \ \\" by (rule_tac x="\ \ \" in exI) auto qed lemma compactin_Union: "\finite \; \S. S \ \ \ compactin X S\ \ compactin X (\\)" by (induction rule: finite_induct) (simp_all add: compactin_Un) lemma compactin_subtopology_imp_compact: assumes "compactin (subtopology X S) K" shows "compactin X K" using assms proof (clarsimp simp add: compactin_def) fix \ define \ where "\ \ (\U. U \ S) ` \" assume "K \ topspace X" and "K \ S" and "\x\\. openin X x" and "K \ \\" then have "\V \ \. openin (subtopology X S) V" "K \ \\" unfolding \_def by (auto simp: openin_subtopology) moreover assume "\\. (\x\\. openin (subtopology X S) x) \ K \ \\ \ (\\. finite \ \ \ \ \ \ K \ \\)" ultimately obtain \ where "finite \" "\ \ \" "K \ \\" by meson then have \: "\U. U \ \ \ V = U \ S" if "V \ \" for V unfolding \_def using that by blast let ?\ = "(\F. @U. U \ \ \ F = U \ S) ` \" show "\\. finite \ \ \ \ \ \ K \ \\" proof (intro exI conjI) show "finite ?\" using \finite \\ by blast show "?\ \ \" using someI_ex [OF \] by blast show "K \ \?\" proof clarsimp fix x assume "x \ K" then show "\V \ \. x \ (SOME U. U \ \ \ V = U \ S)" using \K \ \\\ someI_ex [OF \] by (metis (no_types, lifting) IntD1 Union_iff subsetCE) qed qed qed lemma compact_imp_compactin_subtopology: assumes "compactin X K" "K \ S" shows "compactin (subtopology X S) K" using assms proof (clarsimp simp add: compactin_def) fix \ :: "'a set set" define \ where "\ \ {V. openin X V \ (\U \ \. U = V \ S)}" assume "K \ S" and "K \ topspace X" and "\U\\. openin (subtopology X S) U" and "K \ \\" then have "\V \ \. openin X V" "K \ \\" unfolding \_def by (fastforce simp: subset_eq openin_subtopology)+ moreover assume "\\. (\U\\. openin X U) \ K \ \\ \ (\\. finite \ \ \ \ \ \ K \ \\)" ultimately obtain \ where "finite \" "\ \ \" "K \ \\" by meson let ?\ = "(\F. F \ S) ` \" show "\\. finite \ \ \ \ \ \ K \ \\" proof (intro exI conjI) show "finite ?\" using \finite \\ by blast show "?\ \ \" using \_def \\ \ \\ by blast show "K \ \?\" using \K \ \\\ assms(2) by auto qed qed proposition compact_space_fip: "compact_space X \ (\\. (\C\\. closedin X C) \ (\\. finite \ \ \ \ \ \ \\ \ {}) \ \\ \ {})" (is "_ = ?rhs") proof (cases "topspace X = {}") case True then show ?thesis unfolding compact_space_def by (metis Sup_bot_conv(1) closedin_topspace_empty compactin_empty finite.emptyI finite_UnionD order_refl) next case False show ?thesis proof safe fix \ :: "'a set set" assume * [rule_format]: "\\. finite \ \ \ \ \ \ \\ \ {}" define \ where "\ \ (\S. topspace X - S) ` \" assume clo: "\C\\. closedin X C" and [simp]: "\\ = {}" then have "\V \ \. openin X V" "topspace X \ \\" by (auto simp: \_def) moreover assume [unfolded compact_space_alt, rule_format, of \]: "compact_space X" ultimately obtain \ where \: "finite \" "\ \ \" "topspace X \ topspace X - \\" by (auto simp: ex_finite_subset_image \_def) moreover have "\ \ {}" using \ \topspace X \ {}\ by blast ultimately show "False" using * [of \] by auto (metis Diff_iff Inter_iff clo closedin_def subsetD) next assume R [rule_format]: ?rhs show "compact_space X" unfolding compact_space_alt proof clarify fix \ :: "'a set set" define \ where "\ \ (\S. topspace X - S) ` \" assume "\C\\. openin X C" and "topspace X \ \\" with \topspace X \ {}\ have *: "\V \ \. closedin X V" "\ \ {}" by (auto simp: \_def) show "\\. finite \ \ \ \ \ \ topspace X \ \\" proof (rule ccontr; simp) assume "\\\\. finite \ \ \ topspace X \ \\" then have "\\. finite \ \ \ \ \ \ \\ \ {}" by (simp add: \_def all_finite_subset_image) with \topspace X \ \\\ show False using R [of \] * by (simp add: \_def) qed qed qed qed corollary compactin_fip: "compactin X S \ S \ topspace X \ (\\. (\C\\. closedin X C) \ (\\. finite \ \ \ \ \ \ S \ \\ \ {}) \ S \ \\ \ {})" proof (cases "S = {}") case False show ?thesis proof (cases "S \ topspace X") case True then have "compactin X S \ (\\. \ \ (\T. S \ T) ` {T. closedin X T} \ (\\. finite \ \ \ \ \ \ \\ \ {}) \ \\ \ {})" by (simp add: compact_space_fip compactin_subspace closedin_subtopology image_def subset_eq Int_commute imp_conjL) also have "\ = (\\\Collect (closedin X). (\\. finite \ \ \ \ (\) S ` \ \ \\ \ {}) \ \ ((\) S ` \) \ {})" by (simp add: all_subset_image) also have "\ = (\\. (\C\\. closedin X C) \ (\\. finite \ \ \ \ \ \ S \ \\ \ {}) \ S \ \\ \ {})" proof - have eq: "((\\. finite \ \ \ \ \ \ \ ((\) S ` \) \ {}) \ \ ((\) S ` \) \ {}) \ ((\\. finite \ \ \ \ \ \ S \ \\ \ {}) \ S \ \\ \ {})" for \ by simp (use \S \ {}\ in blast) show ?thesis unfolding imp_conjL [symmetric] all_finite_subset_image eq by blast qed finally show ?thesis using True by simp qed (simp add: compactin_subspace) qed force corollary compact_space_imp_nest: fixes C :: "nat \ 'a set" assumes "compact_space X" and clo: "\n. closedin X (C n)" and ne: "\n. C n \ {}" and inc: "\m n. m \ n \ C n \ C m" shows "(\n. C n) \ {}" proof - let ?\ = "range (\n. \m \ n. C m)" have "closedin X A" if "A \ ?\" for A using that clo by auto moreover have "(\n\K. \m \ n. C m) \ {}" if "finite K" for K proof - obtain n where "\k. k \ K \ k \ n" using Max.coboundedI \finite K\ by blast with inc have "C n \ (\n\K. \m \ n. C m)" by blast with ne [of n] show ?thesis by blast qed ultimately show ?thesis using \compact_space X\ [unfolded compact_space_fip, rule_format, of ?\] by (simp add: all_finite_subset_image INT_extend_simps UN_atMost_UNIV del: INT_simps) qed lemma compactin_discrete_topology: "compactin (discrete_topology X) S \ S \ X \ finite S" (is "?lhs = ?rhs") proof (intro iffI conjI) assume L: ?lhs then show "S \ X" by (auto simp: compactin_def) have *: "\\. Ball \ (openin (discrete_topology X)) \ S \ \\ \ (\\. finite \ \ \ \ \ \ S \ \\)" using L by (auto simp: compactin_def) show "finite S" using * [of "(\x. {x}) ` X"] \S \ X\ by clarsimp (metis UN_singleton finite_subset_image infinite_super) next assume ?rhs then show ?lhs by (simp add: finite_imp_compactin) qed lemma compact_space_discrete_topology: "compact_space(discrete_topology X) \ finite X" by (simp add: compactin_discrete_topology compact_space_def) lemma compact_space_imp_Bolzano_Weierstrass: assumes "compact_space X" "infinite S" "S \ topspace X" shows "X derived_set_of S \ {}" proof assume X: "X derived_set_of S = {}" then have "closedin X S" by (simp add: closedin_contains_derived_set assms) then have "compactin X S" by (rule closedin_compact_space [OF \compact_space X\]) with X show False by (metis \infinite S\ compactin_subspace compact_space_discrete_topology inf_bot_right subtopology_eq_discrete_topology_eq) qed lemma compactin_imp_Bolzano_Weierstrass: "\compactin X S; infinite T \ T \ S\ \ S \ X derived_set_of T \ {}" using compact_space_imp_Bolzano_Weierstrass [of "subtopology X S"] by (simp add: compactin_subspace derived_set_of_subtopology inf_absorb2) lemma compact_closure_of_imp_Bolzano_Weierstrass: "\compactin X (X closure_of S); infinite T; T \ S; T \ topspace X\ \ X derived_set_of T \ {}" using closure_of_mono closure_of_subset compactin_imp_Bolzano_Weierstrass by fastforce lemma discrete_compactin_eq_finite: "S \ X derived_set_of S = {} \ compactin X S \ S \ topspace X \ finite S" by (meson compactin_imp_Bolzano_Weierstrass finite_imp_compactin_eq order_refl) lemma discrete_compact_space_eq_finite: "X derived_set_of (topspace X) = {} \ (compact_space X \ finite(topspace X))" by (metis compact_space_discrete_topology discrete_topology_unique_derived_set) lemma image_compactin: assumes cpt: "compactin X S" and cont: "continuous_map X Y f" shows "compactin Y (f ` S)" unfolding compactin_def proof (intro conjI allI impI) show "f ` S \ topspace Y" using compactin_subset_topspace cont continuous_map_image_subset_topspace cpt by blast next fix \ :: "'b set set" assume \: "Ball \ (openin Y) \ f ` S \ \\" define \ where "\ \ (\U. {x \ topspace X. f x \ U}) ` \" have "S \ topspace X" and *: "\\. \\U\\. openin X U; S \ \\\ \ \\. finite \ \ \ \ \ \ S \ \\" using cpt by (auto simp: compactin_def) obtain \ where \: "finite \" "\ \ \" "S \ \\" proof - have 1: "\U\\. openin X U" unfolding \_def using \ cont[unfolded continuous_map] by blast have 2: "S \ \\" unfolding \_def using compactin_subset_topspace cpt \ by fastforce show thesis using * [OF 1 2] that by metis qed have "\v \ \. \U. U \ \ \ v = {x \ topspace X. f x \ U}" using \_def by blast then obtain U where U: "\v \ \. U v \ \ \ v = {x \ topspace X. f x \ U v}" by metis show "\\. finite \ \ \ \ \ \ f ` S \ \\" proof (intro conjI exI) show "finite (U ` \)" by (simp add: \finite \\) next show "U ` \ \ \" using \\ \ \\ U by auto next show "f ` S \ \ (U ` \)" using \(2-3) U UnionE subset_eq U by fastforce qed qed lemma homeomorphic_compact_space: assumes "X homeomorphic_space Y" shows "compact_space X \ compact_space Y" using homeomorphic_space_sym by (metis assms compact_space_def homeomorphic_eq_everything_map homeomorphic_space image_compactin) lemma homeomorphic_map_compactness: assumes hom: "homeomorphic_map X Y f" and U: "U \ topspace X" shows "compactin Y (f ` U) \ compactin X U" proof - have "f ` U \ topspace Y" using hom U homeomorphic_imp_surjective_map by blast moreover have "homeomorphic_map (subtopology X U) (subtopology Y (f ` U)) f" using U hom homeomorphic_imp_surjective_map by (blast intro: homeomorphic_map_subtopologies) then have "compact_space (subtopology Y (f ` U)) = compact_space (subtopology X U)" using homeomorphic_compact_space homeomorphic_map_imp_homeomorphic_space by blast ultimately show ?thesis by (simp add: compactin_subspace U) qed lemma homeomorphic_map_compactness_eq: "homeomorphic_map X Y f \ compactin X U \ U \ topspace X \ compactin Y (f ` U)" by (meson compactin_subset_topspace homeomorphic_map_compactness) subsection\Embedding maps\ definition embedding_map where "embedding_map X Y f \ homeomorphic_map X (subtopology Y (f ` (topspace X))) f" lemma embedding_map_eq: "\embedding_map X Y f; \x. x \ topspace X \ f x = g x\ \ embedding_map X Y g" unfolding embedding_map_def by (metis homeomorphic_map_eq image_cong) lemma embedding_map_compose: assumes "embedding_map X X' f" "embedding_map X' X'' g" shows "embedding_map X X'' (g \ f)" proof - have hm: "homeomorphic_map X (subtopology X' (f ` topspace X)) f" "homeomorphic_map X' (subtopology X'' (g ` topspace X')) g" using assms by (auto simp: embedding_map_def) then obtain C where "g ` topspace X' \ C = (g \ f) ` topspace X" by (metis (no_types) Int_absorb1 continuous_map_image_subset_topspace continuous_map_in_subtopology homeomorphic_eq_everything_map image_comp image_mono) then have "homeomorphic_map (subtopology X' (f ` topspace X)) (subtopology X'' ((g \ f) ` topspace X)) g" by (metis hm homeomorphic_imp_surjective_map homeomorphic_map_subtopologies image_comp subtopology_subtopology topspace_subtopology) then show ?thesis unfolding embedding_map_def using hm(1) homeomorphic_map_compose by blast qed lemma surjective_embedding_map: "embedding_map X Y f \ f ` (topspace X) = topspace Y \ homeomorphic_map X Y f" by (force simp: embedding_map_def homeomorphic_eq_everything_map) lemma embedding_map_in_subtopology: "embedding_map X (subtopology Y S) f \ embedding_map X Y f \ f ` (topspace X) \ S" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" unfolding embedding_map_def by (metis continuous_map_in_subtopology homeomorphic_imp_continuous_map inf_absorb2 subtopology_subtopology) qed (simp add: embedding_map_def inf.absorb_iff2 subtopology_subtopology) lemma injective_open_imp_embedding_map: "\continuous_map X Y f; open_map X Y f; inj_on f (topspace X)\ \ embedding_map X Y f" unfolding embedding_map_def by (simp add: continuous_map_in_subtopology continuous_open_quotient_map eq_iff homeomorphic_map_def open_map_imp_subset open_map_into_subtopology) lemma injective_closed_imp_embedding_map: "\continuous_map X Y f; closed_map X Y f; inj_on f (topspace X)\ \ embedding_map X Y f" unfolding embedding_map_def by (simp add: closed_map_imp_subset closed_map_into_subtopology continuous_closed_quotient_map continuous_map_in_subtopology dual_order.eq_iff homeomorphic_map_def) lemma embedding_map_imp_homeomorphic_space: "embedding_map X Y f \ X homeomorphic_space (subtopology Y (f ` (topspace X)))" unfolding embedding_map_def using homeomorphic_space by blast lemma embedding_imp_closed_map: "\embedding_map X Y f; closedin Y (f ` topspace X)\ \ closed_map X Y f" unfolding closed_map_def by (auto simp: closedin_closed_subtopology embedding_map_def homeomorphic_map_closedness_eq) subsection\Retraction and section maps\ definition retraction_maps :: "'a topology \ 'b topology \ ('a \ 'b) \ ('b \ 'a) \ bool" where "retraction_maps X Y f g \ continuous_map X Y f \ continuous_map Y X g \ (\x \ topspace Y. f(g x) = x)" definition section_map :: "'a topology \ 'b topology \ ('a \ 'b) \ bool" where "section_map X Y f \ \g. retraction_maps Y X g f" definition retraction_map :: "'a topology \ 'b topology \ ('a \ 'b) \ bool" where "retraction_map X Y f \ \g. retraction_maps X Y f g" lemma retraction_maps_eq: "\retraction_maps X Y f g; \x. x \ topspace X \ f x = f' x; \x. x \ topspace Y \ g x = g' x\ \ retraction_maps X Y f' g'" unfolding retraction_maps_def by (metis (no_types, lifting) continuous_map_def continuous_map_eq) lemma section_map_eq: "\section_map X Y f; \x. x \ topspace X \ f x = g x\ \ section_map X Y g" unfolding section_map_def using retraction_maps_eq by blast lemma retraction_map_eq: "\retraction_map X Y f; \x. x \ topspace X \ f x = g x\ \ retraction_map X Y g" unfolding retraction_map_def using retraction_maps_eq by blast lemma homeomorphic_imp_retraction_maps: "homeomorphic_maps X Y f g \ retraction_maps X Y f g" by (simp add: homeomorphic_maps_def retraction_maps_def) lemma section_and_retraction_eq_homeomorphic_map: "section_map X Y f \ retraction_map X Y f \ homeomorphic_map X Y f" (is "?lhs = ?rhs") proof assume ?lhs then obtain g g' where f: "continuous_map X Y f" and g: "continuous_map Y X g" "\x\topspace X. g (f x) = x" and g': "continuous_map Y X g'" "\x\topspace Y. f (g' x) = x" by (auto simp: retraction_map_def retraction_maps_def section_map_def) then have "homeomorphic_maps X Y f g" by (force simp add: homeomorphic_maps_def continuous_map_def) then show ?rhs using homeomorphic_map_maps by blast next assume ?rhs then show ?lhs unfolding retraction_map_def section_map_def by (meson homeomorphic_imp_retraction_maps homeomorphic_map_maps homeomorphic_maps_sym) qed lemma section_imp_embedding_map: "section_map X Y f \ embedding_map X Y f" unfolding section_map_def embedding_map_def homeomorphic_map_maps retraction_maps_def homeomorphic_maps_def by (force simp: continuous_map_in_subtopology continuous_map_from_subtopology) lemma retraction_imp_quotient_map: assumes "retraction_map X Y f" shows "quotient_map X Y f" unfolding quotient_map_def proof (intro conjI subsetI allI impI) show "f ` topspace X = topspace Y" using assms by (force simp: retraction_map_def retraction_maps_def continuous_map_def) next fix U assume U: "U \ topspace Y" have "openin Y U" if "\x\topspace Y. g x \ topspace X" "\x\topspace Y. f (g x) = x" "openin Y {x \ topspace Y. g x \ {x \ topspace X. f x \ U}}" for g using openin_subopen U that by fastforce then show "openin X {x \ topspace X. f x \ U} = openin Y U" using assms by (auto simp: retraction_map_def retraction_maps_def continuous_map_def) qed lemma retraction_maps_compose: "\retraction_maps X Y f f'; retraction_maps Y Z g g'\ \ retraction_maps X Z (g \ f) (f' \ g')" by (clarsimp simp: retraction_maps_def continuous_map_compose) (simp add: continuous_map_def) lemma retraction_map_compose: "\retraction_map X Y f; retraction_map Y Z g\ \ retraction_map X Z (g \ f)" by (meson retraction_map_def retraction_maps_compose) lemma section_map_compose: "\section_map X Y f; section_map Y Z g\ \ section_map X Z (g \ f)" by (meson retraction_maps_compose section_map_def) lemma surjective_section_eq_homeomorphic_map: "section_map X Y f \ f ` (topspace X) = topspace Y \ homeomorphic_map X Y f" by (meson section_and_retraction_eq_homeomorphic_map section_imp_embedding_map surjective_embedding_map) lemma surjective_retraction_or_section_map: "f ` (topspace X) = topspace Y \ retraction_map X Y f \ section_map X Y f \ retraction_map X Y f" using section_and_retraction_eq_homeomorphic_map surjective_section_eq_homeomorphic_map by fastforce lemma retraction_imp_surjective_map: "retraction_map X Y f \ f ` (topspace X) = topspace Y" by (simp add: retraction_imp_quotient_map quotient_imp_surjective_map) lemma section_imp_injective_map: "\section_map X Y f; x \ topspace X; y \ topspace X\ \ f x = f y \ x = y" by (metis (mono_tags, opaque_lifting) retraction_maps_def section_map_def) lemma retraction_maps_to_retract_maps: "retraction_maps X Y r s \ retraction_maps X (subtopology X (s ` (topspace Y))) (s \ r) id" unfolding retraction_maps_def by (auto simp: continuous_map_compose continuous_map_into_subtopology continuous_map_from_subtopology) subsection \Continuity\ lemma continuous_on_open: "continuous_on S f \ (\T. openin (top_of_set (f ` S)) T \ openin (top_of_set S) (S \ f -` T))" unfolding continuous_on_open_invariant openin_open Int_def vimage_def Int_commute by (simp add: imp_ex imageI conj_commute eq_commute cong: conj_cong) lemma continuous_on_closed: "continuous_on S f \ (\T. closedin (top_of_set (f ` S)) T \ closedin (top_of_set S) (S \ f -` T))" unfolding continuous_on_closed_invariant closedin_closed Int_def vimage_def Int_commute by (simp add: imp_ex imageI conj_commute eq_commute cong: conj_cong) lemma continuous_on_imp_closedin: assumes "continuous_on S f" "closedin (top_of_set (f ` S)) T" shows "closedin (top_of_set S) (S \ f -` T)" using assms continuous_on_closed by blast lemma continuous_map_subtopology_eu [simp]: "continuous_map (top_of_set S) (subtopology euclidean T) h \ continuous_on S h \ h ` S \ T" by (simp add: continuous_map_in_subtopology) lemma continuous_map_euclidean_top_of_set: assumes eq: "f -` S = UNIV" and cont: "continuous_on UNIV f" shows "continuous_map euclidean (top_of_set S) f" by (simp add: cont continuous_map_into_subtopology eq image_subset_iff_subset_vimage) subsection\<^marker>\tag unimportant\ \Half-global and completely global cases\ lemma continuous_openin_preimage_gen: assumes "continuous_on S f" "open T" shows "openin (top_of_set S) (S \ f -` T)" proof - have *: "(S \ f -` T) = (S \ f -` (T \ f ` S))" by auto have "openin (top_of_set (f ` S)) (T \ f ` S)" using openin_open_Int[of T "f ` S", OF assms(2)] unfolding openin_open by auto then show ?thesis using assms(1)[unfolded continuous_on_open, THEN spec[where x="T \ f ` S"]] using * by auto qed lemma continuous_closedin_preimage: assumes "continuous_on S f" and "closed T" shows "closedin (top_of_set S) (S \ f -` T)" proof - have *: "(S \ f -` T) = (S \ f -` (T \ f ` S))" by auto have "closedin (top_of_set (f ` S)) (T \ f ` S)" using closedin_closed_Int[of T "f ` S", OF assms(2)] by (simp add: Int_commute) then show ?thesis using assms(1)[unfolded continuous_on_closed, THEN spec[where x="T \ f ` S"]] using * by auto qed lemma continuous_openin_preimage_eq: "continuous_on S f \ (\T. open T \ openin (top_of_set S) (S \ f -` T))" by (metis Int_commute continuous_on_open_invariant open_openin openin_subtopology) lemma continuous_closedin_preimage_eq: "continuous_on S f \ (\T. closed T \ closedin (top_of_set S) (S \ f -` T))" by (metis Int_commute closedin_closed continuous_on_closed_invariant) lemma continuous_open_preimage: assumes contf: "continuous_on S f" and "open S" "open T" shows "open (S \ f -` T)" proof- obtain U where "open U" "(S \ f -` T) = S \ U" using continuous_openin_preimage_gen[OF contf \open T\] unfolding openin_open by auto then show ?thesis using open_Int[of S U, OF \open S\] by auto qed lemma continuous_closed_preimage: assumes contf: "continuous_on S f" and "closed S" "closed T" shows "closed (S \ f -` T)" proof- obtain U where "closed U" "(S \ f -` T) = S \ U" using continuous_closedin_preimage[OF contf \closed T\] unfolding closedin_closed by auto then show ?thesis using closed_Int[of S U, OF \closed S\] by auto qed lemma continuous_open_vimage: "open S \ (\x. continuous (at x) f) \ open (f -` S)" by (metis continuous_on_eq_continuous_within open_vimage) lemma continuous_closed_vimage: "closed S \ (\x. continuous (at x) f) \ closed (f -` S)" by (simp add: closed_vimage continuous_on_eq_continuous_within) lemma Times_in_interior_subtopology: assumes "(x, y) \ U" "openin (top_of_set (S \ T)) U" obtains V W where "openin (top_of_set S) V" "x \ V" "openin (top_of_set T) W" "y \ W" "(V \ W) \ U" proof - from assms obtain E where "open E" "U = S \ T \ E" "(x, y) \ E" "x \ S" "y \ T" by (auto simp: openin_open) from open_prod_elim[OF \open E\ \(x, y) \ E\] obtain E1 E2 where "open E1" "open E2" "(x, y) \ E1 \ E2" "E1 \ E2 \ E" by blast show ?thesis proof show "openin (top_of_set S) (E1 \ S)" "openin (top_of_set T) (E2 \ T)" using \open E1\ \open E2\ by (auto simp: openin_open) show "x \ E1 \ S" "y \ E2 \ T" using \(x, y) \ E1 \ E2\ \x \ S\ \y \ T\ by auto show "(E1 \ S) \ (E2 \ T) \ U" using \E1 \ E2 \ E\ \U = _\ - by (auto simp: ) + by auto qed qed lemma closedin_Times: "closedin (top_of_set S) S' \ closedin (top_of_set T) T' \ closedin (top_of_set (S \ T)) (S' \ T')" unfolding closedin_closed using closed_Times by blast lemma openin_Times: "openin (top_of_set S) S' \ openin (top_of_set T) T' \ openin (top_of_set (S \ T)) (S' \ T')" unfolding openin_open using open_Times by blast lemma openin_Times_eq: fixes S :: "'a::topological_space set" and T :: "'b::topological_space set" shows "openin (top_of_set (S \ T)) (S' \ T') \ S' = {} \ T' = {} \ openin (top_of_set S) S' \ openin (top_of_set T) T'" (is "?lhs = ?rhs") proof (cases "S' = {} \ T' = {}") case True then show ?thesis by auto next case False then obtain x y where "x \ S'" "y \ T'" by blast show ?thesis proof assume ?lhs have "openin (top_of_set S) S'" proof (subst openin_subopen, clarify) show "\U. openin (top_of_set S) U \ x \ U \ U \ S'" if "x \ S'" for x using that \y \ T'\ Times_in_interior_subtopology [OF _ \?lhs\, of x y] by simp (metis mem_Sigma_iff subsetD subsetI) qed moreover have "openin (top_of_set T) T'" proof (subst openin_subopen, clarify) show "\U. openin (top_of_set T) U \ y \ U \ U \ T'" if "y \ T'" for y using that \x \ S'\ Times_in_interior_subtopology [OF _ \?lhs\, of x y] by simp (metis mem_Sigma_iff subsetD subsetI) qed ultimately show ?rhs by simp next assume ?rhs with False show ?lhs by (simp add: openin_Times) qed qed lemma Lim_transform_within_openin: assumes f: "(f \ l) (at a within T)" and "openin (top_of_set T) S" "a \ S" and eq: "\x. \x \ S; x \ a\ \ f x = g x" shows "(g \ l) (at a within T)" proof - have "\\<^sub>F x in at a within T. x \ T \ x \ a" by (simp add: eventually_at_filter) moreover from \openin _ _\ obtain U where "open U" "S = T \ U" by (auto simp: openin_open) then have "a \ U" using \a \ S\ by auto from topological_tendstoD[OF tendsto_ident_at \open U\ \a \ U\] have "\\<^sub>F x in at a within T. x \ U" by auto ultimately have "\\<^sub>F x in at a within T. f x = g x" by eventually_elim (auto simp: \S = _\ eq) with f show ?thesis by (rule Lim_transform_eventually) qed lemma continuous_on_open_gen: assumes "f ` S \ T" shows "continuous_on S f \ (\U. openin (top_of_set T) U \ openin (top_of_set S) (S \ f -` U))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (clarsimp simp add: continuous_openin_preimage_eq openin_open) (metis Int_assoc assms image_subset_iff_subset_vimage inf.absorb_iff1) next assume R [rule_format]: ?rhs show ?lhs proof (clarsimp simp add: continuous_openin_preimage_eq) fix U::"'a set" assume "open U" then have "openin (top_of_set S) (S \ f -` (U \ T))" by (metis R inf_commute openin_open) then show "openin (top_of_set S) (S \ f -` U)" by (metis Int_assoc Int_commute assms image_subset_iff_subset_vimage inf.absorb_iff2 vimage_Int) qed qed lemma continuous_openin_preimage: "\continuous_on S f; f ` S \ T; openin (top_of_set T) U\ \ openin (top_of_set S) (S \ f -` U)" by (simp add: continuous_on_open_gen) lemma continuous_on_closed_gen: assumes "f ` S \ T" shows "continuous_on S f \ (\U. closedin (top_of_set T) U \ closedin (top_of_set S) (S \ f -` U))" (is "?lhs = ?rhs") proof - have *: "U \ T \ S \ f -` (T - U) = S - (S \ f -` U)" for U using assms by blast show ?thesis proof assume L: ?lhs show ?rhs proof clarify fix U assume "closedin (top_of_set T) U" then show "closedin (top_of_set S) (S \ f -` U)" using L unfolding continuous_on_open_gen [OF assms] by (metis * closedin_def inf_le1 topspace_euclidean_subtopology) qed next assume R [rule_format]: ?rhs show ?lhs unfolding continuous_on_open_gen [OF assms] by (metis * R inf_le1 openin_closedin_eq topspace_euclidean_subtopology) qed qed lemma continuous_closedin_preimage_gen: assumes "continuous_on S f" "f ` S \ T" "closedin (top_of_set T) U" shows "closedin (top_of_set S) (S \ f -` U)" using assms continuous_on_closed_gen by blast lemma continuous_transform_within_openin: assumes "continuous (at a within T) f" and "openin (top_of_set T) S" "a \ S" and eq: "\x. x \ S \ f x = g x" shows "continuous (at a within T) g" using assms by (simp add: Lim_transform_within_openin continuous_within) subsection\<^marker>\tag important\ \The topology generated by some (open) subsets\ text \In the definition below of a generated topology, the \Empty\ case is not necessary, as it follows from \UN\ taking for \K\ the empty set. However, it is convenient to have, and is never a problem in proofs, so I prefer to write it down explicitly. We do not require \UNIV\ to be an open set, as this will not be the case in applications. (We are thinking of a topology on a subset of \UNIV\, the remaining part of \UNIV\ being irrelevant.)\ inductive generate_topology_on for S where Empty: "generate_topology_on S {}" | Int: "generate_topology_on S a \ generate_topology_on S b \ generate_topology_on S (a \ b)" | UN: "(\k. k \ K \ generate_topology_on S k) \ generate_topology_on S (\K)" | Basis: "s \ S \ generate_topology_on S s" lemma istopology_generate_topology_on: "istopology (generate_topology_on S)" unfolding istopology_def by (auto intro: generate_topology_on.intros) text \The basic property of the topology generated by a set \S\ is that it is the smallest topology containing all the elements of \S\:\ lemma generate_topology_on_coarsest: assumes T: "istopology T" "\s. s \ S \ T s" and gen: "generate_topology_on S s0" shows "T s0" using gen by (induct rule: generate_topology_on.induct) (use T in \auto simp: istopology_def\) abbreviation\<^marker>\tag unimportant\ topology_generated_by::"('a set set) \ ('a topology)" where "topology_generated_by S \ topology (generate_topology_on S)" lemma openin_topology_generated_by_iff: "openin (topology_generated_by S) s \ generate_topology_on S s" using topology_inverse'[OF istopology_generate_topology_on[of S]] by simp lemma openin_topology_generated_by: "openin (topology_generated_by S) s \ generate_topology_on S s" using openin_topology_generated_by_iff by auto lemma topology_generated_by_topspace [simp]: "topspace (topology_generated_by S) = (\S)" proof { fix s assume "openin (topology_generated_by S) s" then have "generate_topology_on S s" by (rule openin_topology_generated_by) then have "s \ (\S)" by (induct, auto) } then show "topspace (topology_generated_by S) \ (\S)" unfolding topspace_def by auto next have "generate_topology_on S (\S)" using generate_topology_on.UN[OF generate_topology_on.Basis, of S S] by simp then show "(\S) \ topspace (topology_generated_by S)" unfolding topspace_def using openin_topology_generated_by_iff by auto qed lemma topology_generated_by_Basis: "s \ S \ openin (topology_generated_by S) s" by (simp only: openin_topology_generated_by_iff, auto simp: generate_topology_on.Basis) lemma generate_topology_on_Inter: "\finite \; \K. K \ \ \ generate_topology_on \ K; \ \ {}\ \ generate_topology_on \ (\\)" by (induction \ rule: finite_induct; force intro: generate_topology_on.intros) subsection\Topology bases and sub-bases\ lemma istopology_base_alt: "istopology (arbitrary union_of P) \ (\S T. (arbitrary union_of P) S \ (arbitrary union_of P) T \ (arbitrary union_of P) (S \ T))" by (simp add: istopology_def) (blast intro: arbitrary_union_of_Union) lemma istopology_base_eq: "istopology (arbitrary union_of P) \ (\S T. P S \ P T \ (arbitrary union_of P) (S \ T))" by (simp add: istopology_base_alt arbitrary_union_of_Int_eq) lemma istopology_base: "(\S T. \P S; P T\ \ P(S \ T)) \ istopology (arbitrary union_of P)" by (simp add: arbitrary_def istopology_base_eq union_of_inc) lemma openin_topology_base_unique: "openin X = arbitrary union_of P \ (\V. P V \ openin X V) \ (\U x. openin X U \ x \ U \ (\V. P V \ x \ V \ V \ U))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (auto simp: union_of_def arbitrary_def) next assume R: ?rhs then have *: "\\\Collect P. \\ = S" if "openin X S" for S using that by (rule_tac x="{V. P V \ V \ S}" in exI) fastforce from R show ?lhs by (fastforce simp add: union_of_def arbitrary_def intro: *) qed lemma topology_base_unique: assumes "\S. P S \ openin X S" "\U x. \openin X U; x \ U\ \ \B. P B \ x \ B \ B \ U" shows "topology (arbitrary union_of P) = X" proof - have "X = topology (openin X)" by (simp add: openin_inverse) also from assms have "openin X = arbitrary union_of P" by (subst openin_topology_base_unique) auto finally show ?thesis .. qed lemma topology_bases_eq_aux: "\(arbitrary union_of P) S; \U x. \P U; x \ U\ \ \V. Q V \ x \ V \ V \ U\ \ (arbitrary union_of Q) S" by (metis arbitrary_union_of_alt arbitrary_union_of_idempot) lemma topology_bases_eq: "\\U x. \P U; x \ U\ \ \V. Q V \ x \ V \ V \ U; \V x. \Q V; x \ V\ \ \U. P U \ x \ U \ U \ V\ \ topology (arbitrary union_of P) = topology (arbitrary union_of Q)" by (fastforce intro: arg_cong [where f=topology] elim: topology_bases_eq_aux) lemma istopology_subbase: "istopology (arbitrary union_of (finite intersection_of P relative_to S))" by (simp add: finite_intersection_of_Int istopology_base relative_to_Int) lemma openin_subbase: "openin (topology (arbitrary union_of (finite intersection_of B relative_to U))) S \ (arbitrary union_of (finite intersection_of B relative_to U)) S" by (simp add: istopology_subbase topology_inverse') lemma topspace_subbase [simp]: "topspace(topology (arbitrary union_of (finite intersection_of B relative_to U))) = U" (is "?lhs = _") proof show "?lhs \ U" by (metis arbitrary_union_of_relative_to openin_subbase openin_topspace relative_to_imp_subset) show "U \ ?lhs" by (metis arbitrary_union_of_inc finite_intersection_of_empty inf.orderE istopology_subbase openin_subset relative_to_inc subset_UNIV topology_inverse') qed lemma minimal_topology_subbase: assumes X: "\S. P S \ openin X S" and "openin X U" and S: "openin(topology(arbitrary union_of (finite intersection_of P relative_to U))) S" shows "openin X S" proof - have "(arbitrary union_of (finite intersection_of P relative_to U)) S" using S openin_subbase by blast with X \openin X U\ show ?thesis by (force simp add: union_of_def intersection_of_def relative_to_def intro: openin_Int_Inter) qed lemma istopology_subbase_UNIV: "istopology (arbitrary union_of (finite intersection_of P))" by (simp add: istopology_base finite_intersection_of_Int) lemma generate_topology_on_eq: "generate_topology_on S = arbitrary union_of finite' intersection_of (\x. x \ S)" (is "?lhs = ?rhs") proof (intro ext iffI) fix A assume "?lhs A" then show "?rhs A" proof induction case (Int a b) then show ?case by (metis (mono_tags, lifting) istopology_base_alt finite'_intersection_of_Int istopology_base) next case (UN K) then show ?case by (simp add: arbitrary_union_of_Union) next case (Basis s) then show ?case by (simp add: Sup_upper arbitrary_union_of_inc finite'_intersection_of_inc relative_to_subset) qed auto next fix A assume "?rhs A" then obtain \ where \: "\T. T \ \ \ \\. finite' \ \ \ \ S \ \\ = T" and eq: "A = \\" unfolding union_of_def intersection_of_def by auto show "?lhs A" unfolding eq proof (rule generate_topology_on.UN) fix T assume "T \ \" with \ obtain \ where "finite' \" "\ \ S" "\\ = T" by blast have "generate_topology_on S (\\)" proof (rule generate_topology_on_Inter) show "finite \" "\ \ {}" by (auto simp: \finite' \\) show "\K. K \ \ \ generate_topology_on S K" by (metis \\ \ S\ generate_topology_on.simps subset_iff) qed then show "generate_topology_on S T" using \\\ = T\ by blast qed qed lemma continuous_on_generated_topo_iff: "continuous_map T1 (topology_generated_by S) f \ ((\U. U \ S \ openin T1 (f-`U \ topspace(T1))) \ (f`(topspace T1) \ (\ S)))" unfolding continuous_map_alt topology_generated_by_topspace proof (auto simp add: topology_generated_by_Basis) assume H: "\U. U \ S \ openin T1 (f -` U \ topspace T1)" fix U assume "openin (topology_generated_by S) U" then have "generate_topology_on S U" by (rule openin_topology_generated_by) then show "openin T1 (f -` U \ topspace T1)" proof (induct) fix a b assume H: "openin T1 (f -` a \ topspace T1)" "openin T1 (f -` b \ topspace T1)" have "f -` (a \ b) \ topspace T1 = (f-`a \ topspace T1) \ (f-`b \ topspace T1)" by auto then show "openin T1 (f -` (a \ b) \ topspace T1)" using H by auto next fix K assume H: "openin T1 (f -` k \ topspace T1)" if "k\ K" for k define L where "L = {f -` k \ topspace T1|k. k \ K}" have *: "openin T1 l" if "l \L" for l using that H unfolding L_def by auto have "openin T1 (\L)" using openin_Union[OF *] by simp moreover have "(\L) = (f -` \K \ topspace T1)" unfolding L_def by auto ultimately show "openin T1 (f -` \K \ topspace T1)" by simp qed (auto simp add: H) qed lemma continuous_on_generated_topo: assumes "\U. U \S \ openin T1 (f-`U \ topspace(T1))" "f`(topspace T1) \ (\ S)" shows "continuous_map T1 (topology_generated_by S) f" using assms continuous_on_generated_topo_iff by blast subsection\<^marker>\tag important\ \Pullback topology\ text \Pulling back a topology by map gives again a topology. \subtopology\ is a special case of this notion, pulling back by the identity. We introduce the general notion as we will need it to define the strong operator topology on the space of continuous linear operators, by pulling back the product topology on the space of all functions.\ text \\pullback_topology A f T\ is the pullback of the topology \T\ by the map \f\ on the set \A\.\ definition\<^marker>\tag important\ pullback_topology::"('a set) \ ('a \ 'b) \ ('b topology) \ ('a topology)" where "pullback_topology A f T = topology (\S. \U. openin T U \ S = f-`U \ A)" lemma istopology_pullback_topology: "istopology (\S. \U. openin T U \ S = f-`U \ A)" unfolding istopology_def proof (auto) fix K assume "\S\K. \U. openin T U \ S = f -` U \ A" then have "\U. \S\K. openin T (U S) \ S = f-`(U S) \ A" by (rule bchoice) then obtain U where U: "\S\K. openin T (U S) \ S = f-`(U S) \ A" by blast define V where "V = (\S\K. U S)" have "openin T V" "\K = f -` V \ A" unfolding V_def using U by auto then show "\V. openin T V \ \K = f -` V \ A" by auto qed lemma openin_pullback_topology: "openin (pullback_topology A f T) S \ (\U. openin T U \ S = f-`U \ A)" unfolding pullback_topology_def topology_inverse'[OF istopology_pullback_topology] by auto lemma topspace_pullback_topology: "topspace (pullback_topology A f T) = f-`(topspace T) \ A" by (auto simp add: topspace_def openin_pullback_topology) proposition continuous_map_pullback [intro]: assumes "continuous_map T1 T2 g" shows "continuous_map (pullback_topology A f T1) T2 (g o f)" unfolding continuous_map_alt proof (auto) fix U::"'b set" assume "openin T2 U" then have "openin T1 (g-`U \ topspace T1)" using assms unfolding continuous_map_alt by auto have "(g o f)-`U \ topspace (pullback_topology A f T1) = (g o f)-`U \ A \ f-`(topspace T1)" unfolding topspace_pullback_topology by auto also have "... = f-`(g-`U \ topspace T1) \ A " by auto also have "openin (pullback_topology A f T1) (...)" unfolding openin_pullback_topology using \openin T1 (g-`U \ topspace T1)\ by auto finally show "openin (pullback_topology A f T1) ((g \ f) -` U \ topspace (pullback_topology A f T1))" by auto next fix x assume "x \ topspace (pullback_topology A f T1)" then have "f x \ topspace T1" unfolding topspace_pullback_topology by auto then show "g (f x) \ topspace T2" using assms unfolding continuous_map_def by auto qed proposition continuous_map_pullback' [intro]: assumes "continuous_map T1 T2 (f o g)" "topspace T1 \ g-`A" shows "continuous_map T1 (pullback_topology A f T2) g" unfolding continuous_map_alt proof (auto) fix U assume "openin (pullback_topology A f T2) U" then have "\V. openin T2 V \ U = f-`V \ A" unfolding openin_pullback_topology by auto then obtain V where "openin T2 V" "U = f-`V \ A" by blast then have "g -` U \ topspace T1 = g-`(f-`V \ A) \ topspace T1" by blast also have "... = (f o g)-`V \ (g-`A \ topspace T1)" by auto also have "... = (f o g)-`V \ topspace T1" using assms(2) by auto also have "openin T1 (...)" using assms(1) \openin T2 V\ by auto finally show "openin T1 (g -` U \ topspace T1)" by simp next fix x assume "x \ topspace T1" have "(f o g) x \ topspace T2" using assms(1) \x \ topspace T1\ unfolding continuous_map_def by auto then have "g x \ f-`(topspace T2)" unfolding comp_def by blast moreover have "g x \ A" using assms(2) \x \ topspace T1\ by blast ultimately show "g x \ topspace (pullback_topology A f T2)" unfolding topspace_pullback_topology by blast qed subsection\Proper maps (not a priori assumed continuous) \ definition proper_map where "proper_map X Y f \ closed_map X Y f \ (\y \ topspace Y. compactin X {x \ topspace X. f x = y})" lemma proper_imp_closed_map: "proper_map X Y f \ closed_map X Y f" by (simp add: proper_map_def) lemma proper_map_imp_subset_topspace: "proper_map X Y f \ f ` (topspace X) \ topspace Y" by (simp add: closed_map_imp_subset_topspace proper_map_def) lemma closed_injective_imp_proper_map: assumes f: "closed_map X Y f" and inj: "inj_on f (topspace X)" shows "proper_map X Y f" unfolding proper_map_def proof (clarsimp simp: f) show "compactin X {x \ topspace X. f x = y}" if "y \ topspace Y" for y proof - have "{x \ topspace X. f x = y} = {} \ (\a \ topspace X. {x \ topspace X. f x = y} = {a})" using inj_on_eq_iff [OF inj] by auto then show ?thesis using that by (metis (no_types, lifting) compactin_empty compactin_sing) qed qed lemma injective_imp_proper_eq_closed_map: "inj_on f (topspace X) \ (proper_map X Y f \ closed_map X Y f)" using closed_injective_imp_proper_map proper_imp_closed_map by blast lemma homeomorphic_imp_proper_map: "homeomorphic_map X Y f \ proper_map X Y f" by (simp add: closed_injective_imp_proper_map homeomorphic_eq_everything_map) lemma compactin_proper_map_preimage: assumes f: "proper_map X Y f" and "compactin Y K" shows "compactin X {x. x \ topspace X \ f x \ K}" proof - have "f ` (topspace X) \ topspace Y" by (simp add: f proper_map_imp_subset_topspace) have *: "\y. y \ topspace Y \ compactin X {x \ topspace X. f x = y}" using f by (auto simp: proper_map_def) show ?thesis unfolding compactin_def proof clarsimp show "\\. finite \ \ \ \ \ \ {x \ topspace X. f x \ K} \ \\" if \: "\U\\. openin X U" and sub: "{x \ topspace X. f x \ K} \ \\" for \ proof - have "\y \ K. \\. finite \ \ \ \ \ \ {x \ topspace X. f x = y} \ \\" proof fix y assume "y \ K" then have "compactin X {x \ topspace X. f x = y}" by (metis "*" \compactin Y K\ compactin_subspace subsetD) with \y \ K\ show "\\. finite \ \ \ \ \ \ {x \ topspace X. f x = y} \ \\" unfolding compactin_def using \ sub by fastforce qed then obtain \ where \: "\y. y \ K \ finite (\ y) \ \ y \ \ \ {x \ topspace X. f x = y} \ \(\ y)" by (metis (full_types)) define F where "F \ \y. topspace Y - f ` (topspace X - \(\ y))" have "\\. finite \ \ \ \ F ` K \ K \ \\" proof (rule compactinD [OF \compactin Y K\]) have "\x. x \ K \ closedin Y (f ` (topspace X - \(\ x)))" using f unfolding proper_map_def closed_map_def by (meson \ \ openin_Union openin_closedin_eq subsetD) then show "openin Y U" if "U \ F ` K" for U using that by (auto simp: F_def) show "K \ \(F ` K)" using \ \compactin Y K\ unfolding F_def compactin_def by fastforce qed then obtain J where "finite J" "J \ K" and J: "K \ \(F ` J)" by (auto simp: ex_finite_subset_image) show ?thesis unfolding F_def proof (intro exI conjI) show "finite (\(\ ` J))" using \ \J \ K\ \finite J\ by blast show "\(\ ` J) \ \" using \ \J \ K\ by blast show "{x \ topspace X. f x \ K} \ \(\(\ ` J))" using J \J \ K\ unfolding F_def by auto qed qed qed qed lemma compact_space_proper_map_preimage: assumes f: "proper_map X Y f" and fim: "f ` (topspace X) = topspace Y" and "compact_space Y" shows "compact_space X" proof - have eq: "topspace X = {x \ topspace X. f x \ topspace Y}" using fim by blast moreover have "compactin Y (topspace Y)" using \compact_space Y\ compact_space_def by auto ultimately show ?thesis unfolding compact_space_def using eq f compactin_proper_map_preimage by fastforce qed lemma proper_map_alt: "proper_map X Y f \ closed_map X Y f \ (\K. compactin Y K \ compactin X {x. x \ topspace X \ f x \ K})" proof (intro iffI conjI allI impI) show "compactin X {x \ topspace X. f x \ K}" if "proper_map X Y f" and "compactin Y K" for K using that by (simp add: compactin_proper_map_preimage) show "proper_map X Y f" if f: "closed_map X Y f \ (\K. compactin Y K \ compactin X {x \ topspace X. f x \ K})" proof - have "compactin X {x \ topspace X. f x = y}" if "y \ topspace Y" for y proof - have "compactin X {x \ topspace X. f x \ {y}}" using f compactin_sing that by fastforce then show ?thesis by auto qed with f show ?thesis by (auto simp: proper_map_def) qed qed (simp add: proper_imp_closed_map) lemma proper_map_on_empty: "topspace X = {} \ proper_map X Y f" by (auto simp: proper_map_def closed_map_on_empty) lemma proper_map_id [simp]: "proper_map X X id" proof (clarsimp simp: proper_map_alt closed_map_id) fix K assume K: "compactin X K" then have "{a \ topspace X. a \ K} = K" by (simp add: compactin_subspace subset_antisym subset_iff) then show "compactin X {a \ topspace X. a \ K}" using K by auto qed lemma proper_map_compose: assumes "proper_map X Y f" "proper_map Y Z g" shows "proper_map X Z (g \ f)" proof - have "closed_map X Y f" and f: "\K. compactin Y K \ compactin X {x \ topspace X. f x \ K}" and "closed_map Y Z g" and g: "\K. compactin Z K \ compactin Y {x \ topspace Y. g x \ K}" using assms by (auto simp: proper_map_alt) show ?thesis unfolding proper_map_alt proof (intro conjI allI impI) show "closed_map X Z (g \ f)" using \closed_map X Y f\ \closed_map Y Z g\ closed_map_compose by blast have "{x \ topspace X. g (f x) \ K} = {x \ topspace X. f x \ {b \ topspace Y. g b \ K}}" for K using \closed_map X Y f\ closed_map_imp_subset_topspace by blast then show "compactin X {x \ topspace X. (g \ f) x \ K}" if "compactin Z K" for K using f [OF g [OF that]] by auto qed qed lemma proper_map_const: "proper_map X Y (\x. c) \ compact_space X \ (topspace X = {} \ closedin Y {c})" proof (cases "topspace X = {}") case True then show ?thesis by (simp add: compact_space_topspace_empty proper_map_on_empty) next case False have *: "compactin X {x \ topspace X. c = y}" if "compact_space X" for y proof (cases "c = y") case True then show ?thesis using compact_space_def \compact_space X\ by auto qed auto then show ?thesis using closed_compactin closedin_subset by (force simp: False proper_map_def closed_map_const compact_space_def) qed lemma proper_map_inclusion: "s \ topspace X \ proper_map (subtopology X s) X id \ closedin X s \ (\k. compactin X k \ compactin X (s \ k))" by (auto simp: proper_map_alt closed_map_inclusion_eq inf.absorb_iff2 Collect_conj_eq compactin_subtopology intro: closed_Int_compactin) subsection\Perfect maps (proper, continuous and surjective)\ definition perfect_map where "perfect_map X Y f \ continuous_map X Y f \ proper_map X Y f \ f ` (topspace X) = topspace Y" lemma homeomorphic_imp_perfect_map: "homeomorphic_map X Y f \ perfect_map X Y f" by (simp add: homeomorphic_eq_everything_map homeomorphic_imp_proper_map perfect_map_def) lemma perfect_imp_quotient_map: "perfect_map X Y f \ quotient_map X Y f" by (simp add: continuous_closed_imp_quotient_map perfect_map_def proper_map_def) lemma homeomorphic_eq_injective_perfect_map: "homeomorphic_map X Y f \ perfect_map X Y f \ inj_on f (topspace X)" using homeomorphic_imp_perfect_map homeomorphic_map_def perfect_imp_quotient_map by blast lemma perfect_injective_eq_homeomorphic_map: "perfect_map X Y f \ inj_on f (topspace X) \ homeomorphic_map X Y f" by (simp add: homeomorphic_eq_injective_perfect_map) lemma perfect_map_id [simp]: "perfect_map X X id" by (simp add: homeomorphic_imp_perfect_map) lemma perfect_map_compose: "\perfect_map X Y f; perfect_map Y Z g\ \ perfect_map X Z (g \ f)" by (meson continuous_map_compose perfect_imp_quotient_map perfect_map_def proper_map_compose quotient_map_compose_eq quotient_map_def) lemma perfect_imp_continuous_map: "perfect_map X Y f \ continuous_map X Y f" using perfect_map_def by blast lemma perfect_imp_closed_map: "perfect_map X Y f \ closed_map X Y f" by (simp add: perfect_map_def proper_map_def) lemma perfect_imp_proper_map: "perfect_map X Y f \ proper_map X Y f" by (simp add: perfect_map_def) lemma perfect_imp_surjective_map: "perfect_map X Y f \ f ` (topspace X) = topspace Y" by (simp add: perfect_map_def) end diff --git a/src/HOL/Analysis/Lebesgue_Integral_Substitution.thy b/src/HOL/Analysis/Lebesgue_Integral_Substitution.thy --- a/src/HOL/Analysis/Lebesgue_Integral_Substitution.thy +++ b/src/HOL/Analysis/Lebesgue_Integral_Substitution.thy @@ -1,406 +1,406 @@ (* Title: HOL/Analysis/Lebesgue_Integral_Substitution.thy Author: Manuel Eberl Provides lemmas for integration by substitution for the basic integral types. Note that the substitution function must have a nonnegative derivative. This could probably be weakened somehow. *) section \Integration by Substition for the Lebesgue Integral\ theory Lebesgue_Integral_Substitution imports Interval_Integral begin lemma nn_integral_substitution_aux: fixes f :: "real \ ennreal" assumes Mf: "f \ borel_measurable borel" assumes nonnegf: "\x. f x \ 0" assumes derivg: "\x. x \ {a..b} \ (g has_real_derivative g' x) (at x)" assumes contg': "continuous_on {a..b} g'" assumes derivg_nonneg: "\x. x \ {a..b} \ g' x \ 0" assumes "a < b" shows "(\\<^sup>+x. f x * indicator {g a..g b} x \lborel) = (\\<^sup>+x. f (g x) * g' x * indicator {a..b} x \lborel)" proof- from \a < b\ have [simp]: "a \ b" by simp from derivg have contg: "continuous_on {a..b} g" by (rule has_real_derivative_imp_continuous_on) from this and contg' have Mg: "set_borel_measurable borel {a..b} g" and Mg': "set_borel_measurable borel {a..b} g'" by (simp_all only: set_measurable_continuous_on_ivl) from derivg have derivg': "\x. x \ {a..b} \ (g has_vector_derivative g' x) (at x)" by (simp only: has_field_derivative_iff_has_vector_derivative) have real_ind[simp]: "\A x. enn2real (indicator A x) = indicator A x" by (auto split: split_indicator) have ennreal_ind[simp]: "\A x. ennreal (indicator A x) = indicator A x" by (auto split: split_indicator) have [simp]: "\x A. indicator A (g x) = indicator (g -` A) x" by (auto split: split_indicator) from derivg derivg_nonneg have monog: "\x y. a \ x \ x \ y \ y \ b \ g x \ g y" by (rule deriv_nonneg_imp_mono) simp_all with monog have [simp]: "g a \ g b" by (auto intro: mono_onD) show ?thesis proof (induction rule: borel_measurable_induct[OF Mf, case_names cong set mult add sup]) case (cong f1 f2) from cong.hyps(3) have "f1 = f2" by auto with cong show ?case by simp next case (set A) from set.hyps show ?case proof (induction rule: borel_set_induct) case empty thus ?case by simp next case (interval c d) { fix u v :: real assume asm: "{u..v} \ {g a..g b}" "u \ v" obtain u' v' where u'v': "{a..b} \ g-`{u..v} = {u'..v'}" "u' \ v'" "g u' = u" "g v' = v" using asm by (rule_tac continuous_interval_vimage_Int[OF contg monog, of u v]) simp_all hence "{u'..v'} \ {a..b}" "{u'..v'} \ g -` {u..v}" by blast+ with u'v'(2) have "u' \ g -` {u..v}" "v' \ g -` {u..v}" by auto from u'v'(1) have [simp]: "{a..b} \ g -` {u..v} \ sets borel" by simp have A: "continuous_on {min u' v'..max u' v'} g'" by (simp only: u'v' max_absorb2 min_absorb1) (intro continuous_on_subset[OF contg'], insert u'v', auto) have "\x. x \ {u'..v'} \ (g has_real_derivative g' x) (at x within {u'..v'})" using asm by (intro has_field_derivative_subset[OF derivg] subsetD[OF \{u'..v'} \ {a..b}\]) auto hence B: "\x. min u' v' \ x \ x \ max u' v' \ (g has_vector_derivative g' x) (at x within {min u' v'..max u' v'})" by (simp only: u'v' max_absorb2 min_absorb1) (auto simp: has_field_derivative_iff_has_vector_derivative) have "integrable lborel (\x. indicator ({a..b} \ g -` {u..v}) x *\<^sub>R g' x)" using set_integrable_subset borel_integrable_atLeastAtMost'[OF contg'] by (metis \{u'..v'} \ {a..b}\ eucl_ivals(5) set_integrable_def sets_lborel u'v'(1)) hence "(\\<^sup>+x. ennreal (g' x) * indicator ({a..b} \ g-` {u..v}) x \lborel) = LBINT x:{a..b} \ g-`{u..v}. g' x" unfolding set_lebesgue_integral_def by (subst nn_integral_eq_integral[symmetric]) (auto intro!: derivg_nonneg nn_integral_cong split: split_indicator) also from interval_integral_FTC_finite[OF A B] have "LBINT x:{a..b} \ g-`{u..v}. g' x = v - u" by (simp add: u'v' interval_integral_Icc \u \ v\) finally have "(\\<^sup>+ x. ennreal (g' x) * indicator ({a..b} \ g -` {u..v}) x \lborel) = ennreal (v - u)" . } note A = this have "(\\<^sup>+x. indicator {c..d} (g x) * ennreal (g' x) * indicator {a..b} x \lborel) = (\\<^sup>+ x. ennreal (g' x) * indicator ({a..b} \ g -` {c..d}) x \lborel)" by (intro nn_integral_cong) (simp split: split_indicator) also have "{a..b} \ g-`{c..d} = {a..b} \ g-`{max (g a) c..min (g b) d}" using \a \ b\ \c \ d\ by (auto intro!: monog intro: order.trans) also have "(\\<^sup>+ x. ennreal (g' x) * indicator ... x \lborel) = (if max (g a) c \ min (g b) d then min (g b) d - max (g a) c else 0)" using \c \ d\ by (simp add: A) also have "... = (\\<^sup>+ x. indicator ({g a..g b} \ {c..d}) x \lborel)" by (subst nn_integral_indicator) (auto intro!: measurable_sets Mg simp:) also have "... = (\\<^sup>+ x. indicator {c..d} x * indicator {g a..g b} x \lborel)" by (intro nn_integral_cong) (auto split: split_indicator) finally show ?case .. next case (compl A) note \A \ sets borel\[measurable] from emeasure_mono[of "A \ {g a..g b}" "{g a..g b}" lborel] have [simp]: "emeasure lborel (A \ {g a..g b}) \ top" by (auto simp: top_unique) have [simp]: "g -` A \ {a..b} \ sets borel" by (rule set_borel_measurable_sets[OF Mg]) auto have [simp]: "g -` (-A) \ {a..b} \ sets borel" by (rule set_borel_measurable_sets[OF Mg]) auto have "(\\<^sup>+x. indicator (-A) x * indicator {g a..g b} x \lborel) = (\\<^sup>+x. indicator (-A \ {g a..g b}) x \lborel)" by (rule nn_integral_cong) (simp split: split_indicator) also from compl have "... = emeasure lborel ({g a..g b} - A)" using derivg_nonneg by (simp add: vimage_Compl diff_eq Int_commute[of "-A"]) also have "{g a..g b} - A = {g a..g b} - A \ {g a..g b}" by blast also have "emeasure lborel ... = g b - g a - emeasure lborel (A \ {g a..g b})" - using \A \ sets borel\ by (subst emeasure_Diff) (auto simp: ) + using \A \ sets borel\ by (subst emeasure_Diff) auto also have "emeasure lborel (A \ {g a..g b}) = \\<^sup>+x. indicator A x * indicator {g a..g b} x \lborel" using \A \ sets borel\ by (subst nn_integral_indicator[symmetric], simp, intro nn_integral_cong) (simp split: split_indicator) also have "... = \\<^sup>+ x. indicator (g-`A \ {a..b}) x * ennreal (g' x * indicator {a..b} x) \lborel" (is "_ = ?I") by (subst compl.IH, intro nn_integral_cong) (simp split: split_indicator) also have "g b - g a = LBINT x:{a..b}. g' x" using derivg' unfolding set_lebesgue_integral_def by (intro integral_FTC_atLeastAtMost[symmetric]) (auto intro: continuous_on_subset[OF contg'] has_field_derivative_subset[OF derivg] has_vector_derivative_at_within) also have "ennreal ... = \\<^sup>+ x. g' x * indicator {a..b} x \lborel" using borel_integrable_atLeastAtMost'[OF contg'] unfolding set_lebesgue_integral_def by (subst nn_integral_eq_integral) (simp_all add: mult.commute derivg_nonneg set_integrable_def split: split_indicator) also have Mg'': "(\x. indicator (g -` A \ {a..b}) x * ennreal (g' x * indicator {a..b} x)) \ borel_measurable borel" using Mg' by (intro borel_measurable_times_ennreal borel_measurable_indicator) (simp_all add: mult.commute set_borel_measurable_def) have le: "(\\<^sup>+x. indicator (g-`A \ {a..b}) x * ennreal (g' x * indicator {a..b} x) \lborel) \ (\\<^sup>+x. ennreal (g' x) * indicator {a..b} x \lborel)" by (intro nn_integral_mono) (simp split: split_indicator add: derivg_nonneg) note integrable = borel_integrable_atLeastAtMost'[OF contg'] with le have notinf: "(\\<^sup>+x. indicator (g-`A \ {a..b}) x * ennreal (g' x * indicator {a..b} x) \lborel) \ top" by (auto simp: real_integrable_def nn_integral_set_ennreal mult.commute top_unique set_integrable_def) have "(\\<^sup>+ x. g' x * indicator {a..b} x \lborel) - ?I = \\<^sup>+ x. ennreal (g' x * indicator {a..b} x) - indicator (g -` A \ {a..b}) x * ennreal (g' x * indicator {a..b} x) \lborel" apply (intro nn_integral_diff[symmetric]) apply (insert Mg', simp add: mult.commute set_borel_measurable_def) [] apply (insert Mg'', simp) [] apply (simp split: split_indicator add: derivg_nonneg) apply (rule notinf) apply (simp split: split_indicator add: derivg_nonneg) done also have "... = \\<^sup>+ x. indicator (-A) (g x) * ennreal (g' x) * indicator {a..b} x \lborel" by (intro nn_integral_cong) (simp split: split_indicator) finally show ?case . next case (union f) then have [simp]: "\i. {a..b} \ g -` f i \ sets borel" by (subst Int_commute, intro set_borel_measurable_sets[OF Mg]) auto have "g -` (\i. f i) \ {a..b} = (\i. {a..b} \ g -` f i)" by auto hence "g -` (\i. f i) \ {a..b} \ sets borel" by (auto simp del: UN_simps) have "(\\<^sup>+x. indicator (\i. f i) x * indicator {g a..g b} x \lborel) = \\<^sup>+x. indicator (\i. {g a..g b} \ f i) x \lborel" by (intro nn_integral_cong) (simp split: split_indicator) also from union have "... = emeasure lborel (\i. {g a..g b} \ f i)" by simp also from union have "... = (\i. emeasure lborel ({g a..g b} \ f i))" by (intro suminf_emeasure[symmetric]) (auto simp: disjoint_family_on_def) also from union have "... = (\i. \\<^sup>+x. indicator ({g a..g b} \ f i) x \lborel)" by simp also have "(\i. \\<^sup>+x. indicator ({g a..g b} \ f i) x \lborel) = (\i. \\<^sup>+x. indicator (f i) x * indicator {g a..g b} x \lborel)" by (intro ext nn_integral_cong) (simp split: split_indicator) also from union.IH have "(\i. \\<^sup>+x. indicator (f i) x * indicator {g a..g b} x \lborel) = (\i. \\<^sup>+ x. indicator (f i) (g x) * ennreal (g' x) * indicator {a..b} x \lborel)" by simp also have "(\i. \\<^sup>+ x. indicator (f i) (g x) * ennreal (g' x) * indicator {a..b} x \lborel) = (\i. \\<^sup>+ x. ennreal (g' x * indicator {a..b} x) * indicator ({a..b} \ g -` f i) x \lborel)" by (intro ext nn_integral_cong) (simp split: split_indicator) also have "(\i. ... i) = \\<^sup>+ x. (\i. ennreal (g' x * indicator {a..b} x) * indicator ({a..b} \ g -` f i) x) \lborel" using Mg' apply (intro nn_integral_suminf[symmetric]) apply (rule borel_measurable_times_ennreal, simp add: mult.commute set_borel_measurable_def) apply (rule borel_measurable_indicator, subst sets_lborel) apply (simp_all split: split_indicator add: derivg_nonneg) done also have "(\x i. ennreal (g' x * indicator {a..b} x) * indicator ({a..b} \ g -` f i) x) = (\x i. ennreal (g' x * indicator {a..b} x) * indicator (g -` f i) x)" by (intro ext) (simp split: split_indicator) also have "(\\<^sup>+ x. (\i. ennreal (g' x * indicator {a..b} x) * indicator (g -` f i) x) \lborel) = \\<^sup>+ x. ennreal (g' x * indicator {a..b} x) * (\i. indicator (g -` f i) x) \lborel" by (intro nn_integral_cong) (auto split: split_indicator simp: derivg_nonneg) also from union have "(\x. \i. indicator (g -` f i) x :: ennreal) = (\x. indicator (\i. g -` f i) x)" by (intro ext suminf_indicator) (auto simp: disjoint_family_on_def) also have "(\\<^sup>+x. ennreal (g' x * indicator {a..b} x) * ... x \lborel) = (\\<^sup>+x. indicator (\i. f i) (g x) * ennreal (g' x) * indicator {a..b} x \lborel)" by (intro nn_integral_cong) (simp split: split_indicator) finally show ?case . qed next case (mult f c) note Mf[measurable] = \f \ borel_measurable borel\ let ?I = "indicator {a..b}" have "(\x. f (g x * ?I x) * ennreal (g' x * ?I x)) \ borel_measurable borel" using Mg Mg' by (intro borel_measurable_times_ennreal measurable_compose[OF _ Mf]) (simp_all add: mult.commute set_borel_measurable_def) also have "(\x. f (g x * ?I x) * ennreal (g' x * ?I x)) = (\x. f (g x) * ennreal (g' x) * ?I x)" by (intro ext) (simp split: split_indicator) finally have Mf': "(\x. f (g x) * ennreal (g' x) * ?I x) \ borel_measurable borel" . with mult show ?case by (subst (1 2 3) mult_ac, subst (1 2) nn_integral_cmult) (simp_all add: mult_ac) next case (add f2 f1) let ?I = "indicator {a..b}" { fix f :: "real \ ennreal" assume Mf: "f \ borel_measurable borel" have "(\x. f (g x * ?I x) * ennreal (g' x * ?I x)) \ borel_measurable borel" using Mg Mg' by (intro borel_measurable_times_ennreal measurable_compose[OF _ Mf]) (simp_all add: mult.commute set_borel_measurable_def) also have "(\x. f (g x * ?I x) * ennreal (g' x * ?I x)) = (\x. f (g x) * ennreal (g' x) * ?I x)" by (intro ext) (simp split: split_indicator) finally have "(\x. f (g x) * ennreal (g' x) * ?I x) \ borel_measurable borel" . } note Mf' = this[OF \f1 \ borel_measurable borel\] this[OF \f2 \ borel_measurable borel\] have "(\\<^sup>+ x. (f1 x + f2 x) * indicator {g a..g b} x \lborel) = (\\<^sup>+ x. f1 x * indicator {g a..g b} x + f2 x * indicator {g a..g b} x \lborel)" by (intro nn_integral_cong) (simp split: split_indicator) also from add have "... = (\\<^sup>+ x. f1 (g x) * ennreal (g' x) * indicator {a..b} x \lborel) + (\\<^sup>+ x. f2 (g x) * ennreal (g' x) * indicator {a..b} x \lborel)" by (simp_all add: nn_integral_add) also from add have "... = (\\<^sup>+ x. f1 (g x) * ennreal (g' x) * indicator {a..b} x + f2 (g x) * ennreal (g' x) * indicator {a..b} x \lborel)" by (intro nn_integral_add[symmetric]) (auto simp add: Mf' derivg_nonneg split: split_indicator) also have "... = \\<^sup>+ x. (f1 (g x) + f2 (g x)) * ennreal (g' x) * indicator {a..b} x \lborel" by (intro nn_integral_cong) (simp split: split_indicator add: distrib_right) finally show ?case . next case (sup F) { fix i let ?I = "indicator {a..b}" have "(\x. F i (g x * ?I x) * ennreal (g' x * ?I x)) \ borel_measurable borel" using Mg Mg' by (rule_tac borel_measurable_times_ennreal, rule_tac measurable_compose[OF _ sup.hyps(1)]) (simp_all add: mult.commute set_borel_measurable_def) also have "(\x. F i (g x * ?I x) * ennreal (g' x * ?I x)) = (\x. F i (g x) * ennreal (g' x) * ?I x)" by (intro ext) (simp split: split_indicator) finally have "... \ borel_measurable borel" . } note Mf' = this have "(\\<^sup>+x. (SUP i. F i x) * indicator {g a..g b} x \lborel) = \\<^sup>+x. (SUP i. F i x* indicator {g a..g b} x) \lborel" by (intro nn_integral_cong) (simp split: split_indicator) also from sup have "... = (SUP i. \\<^sup>+x. F i x* indicator {g a..g b} x \lborel)" by (intro nn_integral_monotone_convergence_SUP) (auto simp: incseq_def le_fun_def split: split_indicator) also from sup have "... = (SUP i. \\<^sup>+x. F i (g x) * ennreal (g' x) * indicator {a..b} x \lborel)" by simp also from sup have "... = \\<^sup>+x. (SUP i. F i (g x) * ennreal (g' x) * indicator {a..b} x) \lborel" by (intro nn_integral_monotone_convergence_SUP[symmetric]) (auto simp: incseq_def le_fun_def derivg_nonneg Mf' split: split_indicator intro!: mult_right_mono) also from sup have "... = \\<^sup>+x. (SUP i. F i (g x)) * ennreal (g' x) * indicator {a..b} x \lborel" by (subst mult.assoc, subst mult.commute, subst SUP_mult_left_ennreal) (auto split: split_indicator simp: derivg_nonneg mult_ac) finally show ?case by (simp add: image_comp) qed qed theorem nn_integral_substitution: fixes f :: "real \ real" assumes Mf[measurable]: "set_borel_measurable borel {g a..g b} f" assumes derivg: "\x. x \ {a..b} \ (g has_real_derivative g' x) (at x)" assumes contg': "continuous_on {a..b} g'" assumes derivg_nonneg: "\x. x \ {a..b} \ g' x \ 0" assumes "a \ b" shows "(\\<^sup>+x. f x * indicator {g a..g b} x \lborel) = (\\<^sup>+x. f (g x) * g' x * indicator {a..b} x \lborel)" proof (cases "a = b") assume "a \ b" with \a \ b\ have "a < b" by auto let ?f' = "\x. f x * indicator {g a..g b} x" from derivg derivg_nonneg have monog: "\x y. a \ x \ x \ y \ y \ b \ g x \ g y" by (rule deriv_nonneg_imp_mono) simp_all have bounds: "\x. x \ a \ x \ b \ g x \ g a" "\x. x \ a \ x \ b \ g x \ g b" by (auto intro: monog) from derivg_nonneg have nonneg: "\f x. x \ a \ x \ b \ g' x \ 0 \ f x * ennreal (g' x) \ 0 \ f x \ 0" by (force simp: field_simps) have nonneg': "\x. a \ x \ x \ b \ \ 0 \ f (g x) \ 0 \ f (g x) * g' x \ g' x = 0" by (metis atLeastAtMost_iff derivg_nonneg eq_iff mult_eq_0_iff mult_le_0_iff) have "(\\<^sup>+x. f x * indicator {g a..g b} x \lborel) = (\\<^sup>+x. ennreal (?f' x) * indicator {g a..g b} x \lborel)" by (intro nn_integral_cong) (auto split: split_indicator split_max simp: zero_ennreal.rep_eq ennreal_neg) also have "... = \\<^sup>+ x. ?f' (g x) * ennreal (g' x) * indicator {a..b} x \lborel" using Mf by (subst nn_integral_substitution_aux[OF _ _ derivg contg' derivg_nonneg \a < b\]) (auto simp add: mult.commute set_borel_measurable_def) also have "... = \\<^sup>+ x. f (g x) * ennreal (g' x) * indicator {a..b} x \lborel" by (intro nn_integral_cong) (auto split: split_indicator simp: max_def dest: bounds) also have "... = \\<^sup>+x. ennreal (f (g x) * g' x * indicator {a..b} x) \lborel" by (intro nn_integral_cong) (auto simp: mult.commute derivg_nonneg ennreal_mult' split: split_indicator) finally show ?thesis . qed auto theorem integral_substitution: assumes integrable: "set_integrable lborel {g a..g b} f" assumes derivg: "\x. x \ {a..b} \ (g has_real_derivative g' x) (at x)" assumes contg': "continuous_on {a..b} g'" assumes derivg_nonneg: "\x. x \ {a..b} \ g' x \ 0" assumes "a \ b" shows "set_integrable lborel {a..b} (\x. f (g x) * g' x)" and "(LBINT x. f x * indicator {g a..g b} x) = (LBINT x. f (g x) * g' x * indicator {a..b} x)" proof- from derivg have contg: "continuous_on {a..b} g" by (rule has_real_derivative_imp_continuous_on) with contg' have Mg: "set_borel_measurable borel {a..b} g" and Mg': "set_borel_measurable borel {a..b} g'" by (simp_all only: set_measurable_continuous_on_ivl) from derivg derivg_nonneg have monog: "\x y. a \ x \ x \ y \ y \ b \ g x \ g y" by (rule deriv_nonneg_imp_mono) simp_all have "(\x. ennreal (f x) * indicator {g a..g b} x) = (\x. ennreal (f x * indicator {g a..g b} x))" by (intro ext) (simp split: split_indicator) with integrable have M1: "(\x. f x * indicator {g a..g b} x) \ borel_measurable borel" by (force simp: mult.commute set_integrable_def) from integrable have M2: "(\x. -f x * indicator {g a..g b} x) \ borel_measurable borel" by (force simp: mult.commute set_integrable_def) have "LBINT x. (f x :: real) * indicator {g a..g b} x = enn2real (\\<^sup>+ x. ennreal (f x) * indicator {g a..g b} x \lborel) - enn2real (\\<^sup>+ x. ennreal (- (f x)) * indicator {g a..g b} x \lborel)" using integrable unfolding set_integrable_def by (subst real_lebesgue_integral_def) (simp_all add: nn_integral_set_ennreal mult.commute) also have *: "(\\<^sup>+x. ennreal (f x) * indicator {g a..g b} x \lborel) = (\\<^sup>+x. ennreal (f x * indicator {g a..g b} x) \lborel)" by (intro nn_integral_cong) (simp split: split_indicator) also from M1 * have A: "(\\<^sup>+ x. ennreal (f x * indicator {g a..g b} x) \lborel) = (\\<^sup>+ x. ennreal (f (g x) * g' x * indicator {a..b} x) \lborel)" by (subst nn_integral_substitution[OF _ derivg contg' derivg_nonneg \a \ b\]) (auto simp: nn_integral_set_ennreal mult.commute set_borel_measurable_def) also have **: "(\\<^sup>+ x. ennreal (- (f x)) * indicator {g a..g b} x \lborel) = (\\<^sup>+ x. ennreal (- (f x) * indicator {g a..g b} x) \lborel)" by (intro nn_integral_cong) (simp split: split_indicator) also from M2 ** have B: "(\\<^sup>+ x. ennreal (- (f x) * indicator {g a..g b} x) \lborel) = (\\<^sup>+ x. ennreal (- (f (g x)) * g' x * indicator {a..b} x) \lborel)" by (subst nn_integral_substitution[OF _ derivg contg' derivg_nonneg \a \ b\]) (auto simp: nn_integral_set_ennreal mult.commute set_borel_measurable_def) also { from integrable have Mf: "set_borel_measurable borel {g a..g b} f" unfolding set_borel_measurable_def set_integrable_def by simp from measurable_compose Mg Mf Mg' borel_measurable_times have "(\x. f (g x * indicator {a..b} x) * indicator {g a..g b} (g x * indicator {a..b} x) * (g' x * indicator {a..b} x)) \ borel_measurable borel" (is "?f \ _") by (simp add: mult.commute set_borel_measurable_def) also have "?f = (\x. f (g x) * g' x * indicator {a..b} x)" using monog by (intro ext) (auto split: split_indicator) finally show "set_integrable lborel {a..b} (\x. f (g x) * g' x)" using A B integrable unfolding real_integrable_def set_integrable_def by (simp_all add: nn_integral_set_ennreal mult.commute) } note integrable' = this have "enn2real (\\<^sup>+ x. ennreal (f (g x) * g' x * indicator {a..b} x) \lborel) - enn2real (\\<^sup>+ x. ennreal (-f (g x) * g' x * indicator {a..b} x) \lborel) = (LBINT x. f (g x) * g' x * indicator {a..b} x)" using integrable' unfolding set_integrable_def by (subst real_lebesgue_integral_def) (simp_all add: field_simps) finally show "(LBINT x. f x * indicator {g a..g b} x) = (LBINT x. f (g x) * g' x * indicator {a..b} x)" . qed theorem interval_integral_substitution: assumes integrable: "set_integrable lborel {g a..g b} f" assumes derivg: "\x. x \ {a..b} \ (g has_real_derivative g' x) (at x)" assumes contg': "continuous_on {a..b} g'" assumes derivg_nonneg: "\x. x \ {a..b} \ g' x \ 0" assumes "a \ b" shows "set_integrable lborel {a..b} (\x. f (g x) * g' x)" and "(LBINT x=g a..g b. f x) = (LBINT x=a..b. f (g x) * g' x)" apply (rule integral_substitution[OF assms], simp, simp) apply (subst (1 2) interval_integral_Icc, fact) apply (rule deriv_nonneg_imp_mono[OF derivg derivg_nonneg], simp, simp, fact) using integral_substitution(2)[OF assms] apply (simp add: mult.commute set_lebesgue_integral_def) done lemma set_borel_integrable_singleton[simp]: "set_integrable lborel {x} (f :: real \ real)" unfolding set_integrable_def by (subst integrable_discrete_difference[where X="{x}" and g="\_. 0"]) auto end diff --git a/src/HOL/Analysis/Linear_Algebra.thy b/src/HOL/Analysis/Linear_Algebra.thy --- a/src/HOL/Analysis/Linear_Algebra.thy +++ b/src/HOL/Analysis/Linear_Algebra.thy @@ -1,1916 +1,1916 @@ (* Title: HOL/Analysis/Linear_Algebra.thy Author: Amine Chaieb, University of Cambridge *) section \Elementary Linear Algebra on Euclidean Spaces\ theory Linear_Algebra imports Euclidean_Space "HOL-Library.Infinite_Set" begin lemma linear_simps: assumes "bounded_linear f" shows "f (a + b) = f a + f b" "f (a - b) = f a - f b" "f 0 = 0" "f (- a) = - f a" "f (s *\<^sub>R v) = s *\<^sub>R (f v)" proof - interpret f: bounded_linear f by fact show "f (a + b) = f a + f b" by (rule f.add) show "f (a - b) = f a - f b" by (rule f.diff) show "f 0 = 0" by (rule f.zero) show "f (- a) = - f a" by (rule f.neg) show "f (s *\<^sub>R v) = s *\<^sub>R (f v)" by (rule f.scale) qed lemma finite_Atleast_Atmost_nat[simp]: "finite {f x |x. x \ (UNIV::'a::finite set)}" using finite finite_image_set by blast lemma substdbasis_expansion_unique: includes inner_syntax assumes d: "d \ Basis" shows "(\i\d. f i *\<^sub>R i) = (x::'a::euclidean_space) \ (\i\Basis. (i \ d \ f i = x \ i) \ (i \ d \ x \ i = 0))" proof - have *: "\x a b P. x * (if P then a else b) = (if P then x * a else x * b)" by auto have **: "finite d" by (auto intro: finite_subset[OF assms]) have ***: "\i. i \ Basis \ (\i\d. f i *\<^sub>R i) \ i = (\x\d. if x = i then f x else 0)" using d by (auto intro!: sum.cong simp: inner_Basis inner_sum_left) show ?thesis unfolding euclidean_eq_iff[where 'a='a] by (auto simp: sum.delta[OF **] ***) qed lemma independent_substdbasis: "d \ Basis \ independent d" by (rule independent_mono[OF independent_Basis]) lemma subset_translation_eq [simp]: fixes a :: "'a::real_vector" shows "(+) a ` s \ (+) a ` t \ s \ t" by auto lemma translate_inj_on: fixes A :: "'a::ab_group_add set" shows "inj_on (\x. a + x) A" unfolding inj_on_def by auto lemma translation_assoc: fixes a b :: "'a::ab_group_add" shows "(\x. b + x) ` ((\x. a + x) ` S) = (\x. (a + b) + x) ` S" by auto lemma translation_invert: fixes a :: "'a::ab_group_add" assumes "(\x. a + x) ` A = (\x. a + x) ` B" shows "A = B" proof - have "(\x. -a + x) ` ((\x. a + x) ` A) = (\x. - a + x) ` ((\x. a + x) ` B)" using assms by auto then show ?thesis using translation_assoc[of "-a" a A] translation_assoc[of "-a" a B] by auto qed lemma translation_galois: fixes a :: "'a::ab_group_add" shows "T = ((\x. a + x) ` S) \ S = ((\x. (- a) + x) ` T)" using translation_assoc[of "-a" a S] apply auto using translation_assoc[of a "-a" T] apply auto done lemma translation_inverse_subset: assumes "((\x. - a + x) ` V) \ (S :: 'n::ab_group_add set)" shows "V \ ((\x. a + x) ` S)" proof - { fix x assume "x \ V" then have "x-a \ S" using assms by auto then have "x \ {a + v |v. v \ S}" apply auto apply (rule exI[of _ "x-a"], simp) done then have "x \ ((\x. a+x) ` S)" by auto } then show ?thesis by auto qed subsection\<^marker>\tag unimportant\ \More interesting properties of the norm\ unbundle inner_syntax text\Equality of vectors in terms of \<^term>\(\)\ products.\ lemma linear_componentwise: fixes f:: "'a::euclidean_space \ 'b::real_inner" assumes lf: "linear f" shows "(f x) \ j = (\i\Basis. (x\i) * (f i\j))" (is "?lhs = ?rhs") proof - interpret linear f by fact have "?rhs = (\i\Basis. (x\i) *\<^sub>R (f i))\j" by (simp add: inner_sum_left) then show ?thesis by (simp add: euclidean_representation sum[symmetric] scale[symmetric]) qed lemma vector_eq: "x = y \ x \ x = x \ y \ y \ y = x \ x" (is "?lhs \ ?rhs") proof assume ?lhs then show ?rhs by simp next assume ?rhs then have "x \ x - x \ y = 0 \ x \ y - y \ y = 0" by simp then have "x \ (x - y) = 0 \ y \ (x - y) = 0" by (simp add: inner_diff inner_commute) then have "(x - y) \ (x - y) = 0" by (simp add: field_simps inner_diff inner_commute) then show "x = y" by simp qed lemma norm_triangle_half_r: "norm (y - x1) < e / 2 \ norm (y - x2) < e / 2 \ norm (x1 - x2) < e" using dist_triangle_half_r unfolding dist_norm[symmetric] by auto lemma norm_triangle_half_l: assumes "norm (x - y) < e / 2" and "norm (x' - y) < e / 2" shows "norm (x - x') < e" using dist_triangle_half_l[OF assms[unfolded dist_norm[symmetric]]] unfolding dist_norm[symmetric] . lemma abs_triangle_half_r: fixes y :: "'a::linordered_field" shows "abs (y - x1) < e / 2 \ abs (y - x2) < e / 2 \ abs (x1 - x2) < e" by linarith lemma abs_triangle_half_l: fixes y :: "'a::linordered_field" assumes "abs (x - y) < e / 2" and "abs (x' - y) < e / 2" shows "abs (x - x') < e" using assms by linarith lemma sum_clauses: shows "sum f {} = 0" and "finite S \ sum f (insert x S) = (if x \ S then sum f S else f x + sum f S)" by (auto simp add: insert_absorb) lemma vector_eq_ldot: "(\x. x \ y = x \ z) \ y = z" proof assume "\x. x \ y = x \ z" then have "\x. x \ (y - z) = 0" by (simp add: inner_diff) then have "(y - z) \ (y - z) = 0" .. then show "y = z" by simp qed simp lemma vector_eq_rdot: "(\z. x \ z = y \ z) \ x = y" proof assume "\z. x \ z = y \ z" then have "\z. (x - y) \ z = 0" by (simp add: inner_diff) then have "(x - y) \ (x - y) = 0" .. then show "x = y" by simp qed simp subsection \Substandard Basis\ lemma ex_card: assumes "n \ card A" shows "\S\A. card S = n" proof (cases "finite A") case True from ex_bij_betw_nat_finite[OF this] obtain f where f: "bij_betw f {0..n \ card A\ have "{..< n} \ {..< card A}" "inj_on f {..< n}" by (auto simp: bij_betw_def intro: subset_inj_on) ultimately have "f ` {..< n} \ A" "card (f ` {..< n}) = n" by (auto simp: bij_betw_def card_image) then show ?thesis by blast next case False with \n \ card A\ show ?thesis by force qed lemma subspace_substandard: "subspace {x::'a::euclidean_space. (\i\Basis. P i \ x\i = 0)}" by (auto simp: subspace_def inner_add_left) lemma dim_substandard: assumes d: "d \ Basis" shows "dim {x::'a::euclidean_space. \i\Basis. i \ d \ x\i = 0} = card d" (is "dim ?A = _") proof (rule dim_unique) from d show "d \ ?A" by (auto simp: inner_Basis) from d show "independent d" by (rule independent_mono [OF independent_Basis]) have "x \ span d" if "\i\Basis. i \ d \ x \ i = 0" for x proof - have "finite d" by (rule finite_subset [OF d finite_Basis]) then have "(\i\d. (x \ i) *\<^sub>R i) \ span d" by (simp add: span_sum span_clauses) also have "(\i\d. (x \ i) *\<^sub>R i) = (\i\Basis. (x \ i) *\<^sub>R i)" by (rule sum.mono_neutral_cong_left [OF finite_Basis d]) (auto simp: that) finally show "x \ span d" by (simp only: euclidean_representation) qed then show "?A \ span d" by auto qed simp subsection \Orthogonality\ definition\<^marker>\tag important\ (in real_inner) "orthogonal x y \ x \ y = 0" context real_inner begin lemma orthogonal_self: "orthogonal x x \ x = 0" by (simp add: orthogonal_def) lemma orthogonal_clauses: "orthogonal a 0" "orthogonal a x \ orthogonal a (c *\<^sub>R x)" "orthogonal a x \ orthogonal a (- x)" "orthogonal a x \ orthogonal a y \ orthogonal a (x + y)" "orthogonal a x \ orthogonal a y \ orthogonal a (x - y)" "orthogonal 0 a" "orthogonal x a \ orthogonal (c *\<^sub>R x) a" "orthogonal x a \ orthogonal (- x) a" "orthogonal x a \ orthogonal y a \ orthogonal (x + y) a" "orthogonal x a \ orthogonal y a \ orthogonal (x - y) a" unfolding orthogonal_def inner_add inner_diff by auto end lemma orthogonal_commute: "orthogonal x y \ orthogonal y x" by (simp add: orthogonal_def inner_commute) lemma orthogonal_scaleR [simp]: "c \ 0 \ orthogonal (c *\<^sub>R x) = orthogonal x" by (rule ext) (simp add: orthogonal_def) lemma pairwise_ortho_scaleR: "pairwise (\i j. orthogonal (f i) (g j)) B \ pairwise (\i j. orthogonal (a i *\<^sub>R f i) (a j *\<^sub>R g j)) B" by (auto simp: pairwise_def orthogonal_clauses) lemma orthogonal_rvsum: "\finite s; \y. y \ s \ orthogonal x (f y)\ \ orthogonal x (sum f s)" by (induction s rule: finite_induct) (auto simp: orthogonal_clauses) lemma orthogonal_lvsum: "\finite s; \x. x \ s \ orthogonal (f x) y\ \ orthogonal (sum f s) y" by (induction s rule: finite_induct) (auto simp: orthogonal_clauses) lemma norm_add_Pythagorean: assumes "orthogonal a b" shows "norm(a + b) ^ 2 = norm a ^ 2 + norm b ^ 2" proof - from assms have "(a - (0 - b)) \ (a - (0 - b)) = a \ a - (0 - b \ b)" by (simp add: algebra_simps orthogonal_def inner_commute) then show ?thesis by (simp add: power2_norm_eq_inner) qed lemma norm_sum_Pythagorean: assumes "finite I" "pairwise (\i j. orthogonal (f i) (f j)) I" shows "(norm (sum f I))\<^sup>2 = (\i\I. (norm (f i))\<^sup>2)" using assms proof (induction I rule: finite_induct) case empty then show ?case by simp next case (insert x I) then have "orthogonal (f x) (sum f I)" by (metis pairwise_insert orthogonal_rvsum) with insert show ?case by (simp add: pairwise_insert norm_add_Pythagorean) qed subsection \Orthogonality of a transformation\ definition\<^marker>\tag important\ "orthogonal_transformation f \ linear f \ (\v w. f v \ f w = v \ w)" lemma\<^marker>\tag unimportant\ orthogonal_transformation: "orthogonal_transformation f \ linear f \ (\v. norm (f v) = norm v)" unfolding orthogonal_transformation_def apply auto apply (erule_tac x=v in allE)+ apply (simp add: norm_eq_sqrt_inner) apply (simp add: dot_norm linear_add[symmetric]) done lemma\<^marker>\tag unimportant\ orthogonal_transformation_id [simp]: "orthogonal_transformation (\x. x)" by (simp add: linear_iff orthogonal_transformation_def) lemma\<^marker>\tag unimportant\ orthogonal_orthogonal_transformation: "orthogonal_transformation f \ orthogonal (f x) (f y) \ orthogonal x y" by (simp add: orthogonal_def orthogonal_transformation_def) lemma\<^marker>\tag unimportant\ orthogonal_transformation_compose: "\orthogonal_transformation f; orthogonal_transformation g\ \ orthogonal_transformation(f \ g)" by (auto simp: orthogonal_transformation_def linear_compose) lemma\<^marker>\tag unimportant\ orthogonal_transformation_neg: "orthogonal_transformation(\x. -(f x)) \ orthogonal_transformation f" by (auto simp: orthogonal_transformation_def dest: linear_compose_neg) lemma\<^marker>\tag unimportant\ orthogonal_transformation_scaleR: "orthogonal_transformation f \ f (c *\<^sub>R v) = c *\<^sub>R f v" by (simp add: linear_iff orthogonal_transformation_def) lemma\<^marker>\tag unimportant\ orthogonal_transformation_linear: "orthogonal_transformation f \ linear f" by (simp add: orthogonal_transformation_def) lemma\<^marker>\tag unimportant\ orthogonal_transformation_inj: "orthogonal_transformation f \ inj f" unfolding orthogonal_transformation_def inj_on_def by (metis vector_eq) lemma\<^marker>\tag unimportant\ orthogonal_transformation_surj: "orthogonal_transformation f \ surj f" for f :: "'a::euclidean_space \ 'a::euclidean_space" by (simp add: linear_injective_imp_surjective orthogonal_transformation_inj orthogonal_transformation_linear) lemma\<^marker>\tag unimportant\ orthogonal_transformation_bij: "orthogonal_transformation f \ bij f" for f :: "'a::euclidean_space \ 'a::euclidean_space" by (simp add: bij_def orthogonal_transformation_inj orthogonal_transformation_surj) lemma\<^marker>\tag unimportant\ orthogonal_transformation_inv: "orthogonal_transformation f \ orthogonal_transformation (inv f)" for f :: "'a::euclidean_space \ 'a::euclidean_space" by (metis (no_types, opaque_lifting) bijection.inv_right bijection_def inj_linear_imp_inv_linear orthogonal_transformation orthogonal_transformation_bij orthogonal_transformation_inj) lemma\<^marker>\tag unimportant\ orthogonal_transformation_norm: "orthogonal_transformation f \ norm (f x) = norm x" by (metis orthogonal_transformation) subsection \Bilinear functions\ definition\<^marker>\tag important\ bilinear :: "('a::real_vector \ 'b::real_vector \ 'c::real_vector) \ bool" where "bilinear f \ (\x. linear (\y. f x y)) \ (\y. linear (\x. f x y))" lemma bilinear_ladd: "bilinear h \ h (x + y) z = h x z + h y z" by (simp add: bilinear_def linear_iff) lemma bilinear_radd: "bilinear h \ h x (y + z) = h x y + h x z" by (simp add: bilinear_def linear_iff) lemma bilinear_times: fixes c::"'a::real_algebra" shows "bilinear (\x y::'a. x*y)" by (auto simp: bilinear_def distrib_left distrib_right intro!: linearI) lemma bilinear_lmul: "bilinear h \ h (c *\<^sub>R x) y = c *\<^sub>R h x y" by (simp add: bilinear_def linear_iff) lemma bilinear_rmul: "bilinear h \ h x (c *\<^sub>R y) = c *\<^sub>R h x y" by (simp add: bilinear_def linear_iff) lemma bilinear_lneg: "bilinear h \ h (- x) y = - h x y" by (drule bilinear_lmul [of _ "- 1"]) simp lemma bilinear_rneg: "bilinear h \ h x (- y) = - h x y" by (drule bilinear_rmul [of _ _ "- 1"]) simp lemma (in ab_group_add) eq_add_iff: "x = x + y \ y = 0" using add_left_imp_eq[of x y 0] by auto lemma bilinear_lzero: assumes "bilinear h" shows "h 0 x = 0" using bilinear_ladd [OF assms, of 0 0 x] by (simp add: eq_add_iff field_simps) lemma bilinear_rzero: assumes "bilinear h" shows "h x 0 = 0" using bilinear_radd [OF assms, of x 0 0 ] by (simp add: eq_add_iff field_simps) lemma bilinear_lsub: "bilinear h \ h (x - y) z = h x z - h y z" using bilinear_ladd [of h x "- y"] by (simp add: bilinear_lneg) lemma bilinear_rsub: "bilinear h \ h z (x - y) = h z x - h z y" using bilinear_radd [of h _ x "- y"] by (simp add: bilinear_rneg) lemma bilinear_sum: assumes "bilinear h" shows "h (sum f S) (sum g T) = sum (\(i,j). h (f i) (g j)) (S \ T) " proof - interpret l: linear "\x. h x y" for y using assms by (simp add: bilinear_def) interpret r: linear "\y. h x y" for x using assms by (simp add: bilinear_def) have "h (sum f S) (sum g T) = sum (\x. h (f x) (sum g T)) S" by (simp add: l.sum) also have "\ = sum (\x. sum (\y. h (f x) (g y)) T) S" by (rule sum.cong) (simp_all add: r.sum) finally show ?thesis unfolding sum.cartesian_product . qed subsection \Adjoints\ definition\<^marker>\tag important\ adjoint :: "(('a::real_inner) \ ('b::real_inner)) \ 'b \ 'a" where "adjoint f = (SOME f'. \x y. f x \ y = x \ f' y)" lemma adjoint_unique: assumes "\x y. inner (f x) y = inner x (g y)" shows "adjoint f = g" unfolding adjoint_def proof (rule some_equality) show "\x y. inner (f x) y = inner x (g y)" by (rule assms) next fix h assume "\x y. inner (f x) y = inner x (h y)" then have "\x y. inner x (g y) = inner x (h y)" using assms by simp then have "\x y. inner x (g y - h y) = 0" by (simp add: inner_diff_right) then have "\y. inner (g y - h y) (g y - h y) = 0" by simp then have "\y. h y = g y" by simp then show "h = g" by (simp add: ext) qed text \TODO: The following lemmas about adjoints should hold for any Hilbert space (i.e. complete inner product space). (see \<^url>\https://en.wikipedia.org/wiki/Hermitian_adjoint\) \ lemma adjoint_works: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes lf: "linear f" shows "x \ adjoint f y = f x \ y" proof - interpret linear f by fact have "\y. \w. \x. f x \ y = x \ w" proof (intro allI exI) fix y :: "'m" and x let ?w = "(\i\Basis. (f i \ y) *\<^sub>R i) :: 'n" have "f x \ y = f (\i\Basis. (x \ i) *\<^sub>R i) \ y" by (simp add: euclidean_representation) also have "\ = (\i\Basis. (x \ i) *\<^sub>R f i) \ y" by (simp add: sum scale) finally show "f x \ y = x \ ?w" by (simp add: inner_sum_left inner_sum_right mult.commute) qed then show ?thesis unfolding adjoint_def choice_iff by (intro someI2_ex[where Q="\f'. x \ f' y = f x \ y"]) auto qed lemma adjoint_clauses: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes lf: "linear f" shows "x \ adjoint f y = f x \ y" and "adjoint f y \ x = y \ f x" by (simp_all add: adjoint_works[OF lf] inner_commute) lemma adjoint_linear: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes lf: "linear f" shows "linear (adjoint f)" by (simp add: lf linear_iff euclidean_eq_iff[where 'a='n] euclidean_eq_iff[where 'a='m] adjoint_clauses[OF lf] inner_distrib) lemma adjoint_adjoint: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes lf: "linear f" shows "adjoint (adjoint f) = f" by (rule adjoint_unique, simp add: adjoint_clauses [OF lf]) subsection\<^marker>\tag unimportant\ \Euclidean Spaces as Typeclass\ lemma independent_Basis: "independent Basis" by (rule independent_Basis) lemma span_Basis [simp]: "span Basis = UNIV" by (rule span_Basis) lemma in_span_Basis: "x \ span Basis" unfolding span_Basis .. subsection\<^marker>\tag unimportant\ \Linearity and Bilinearity continued\ lemma linear_bounded: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes lf: "linear f" shows "\B. \x. norm (f x) \ B * norm x" proof interpret linear f by fact let ?B = "\b\Basis. norm (f b)" show "\x. norm (f x) \ ?B * norm x" proof fix x :: 'a let ?g = "\b. (x \ b) *\<^sub>R f b" have "norm (f x) = norm (f (\b\Basis. (x \ b) *\<^sub>R b))" unfolding euclidean_representation .. also have "\ = norm (sum ?g Basis)" by (simp add: sum scale) finally have th0: "norm (f x) = norm (sum ?g Basis)" . have th: "norm (?g i) \ norm (f i) * norm x" if "i \ Basis" for i proof - from Basis_le_norm[OF that, of x] show "norm (?g i) \ norm (f i) * norm x" unfolding norm_scaleR by (metis mult.commute mult_left_mono norm_ge_zero) qed from sum_norm_le[of _ ?g, OF th] show "norm (f x) \ ?B * norm x" unfolding th0 sum_distrib_right by metis qed qed lemma linear_conv_bounded_linear: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" shows "linear f \ bounded_linear f" proof assume "linear f" then interpret f: linear f . show "bounded_linear f" proof have "\B. \x. norm (f x) \ B * norm x" using \linear f\ by (rule linear_bounded) then show "\K. \x. norm (f x) \ norm x * K" by (simp add: mult.commute) qed next assume "bounded_linear f" then interpret f: bounded_linear f . show "linear f" .. qed lemmas linear_linear = linear_conv_bounded_linear[symmetric] lemma inj_linear_imp_inv_bounded_linear: fixes f::"'a::euclidean_space \ 'a" shows "\bounded_linear f; inj f\ \ bounded_linear (inv f)" by (simp add: inj_linear_imp_inv_linear linear_linear) lemma linear_bounded_pos: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes lf: "linear f" obtains B where "B > 0" "\x. norm (f x) \ B * norm x" proof - have "\B > 0. \x. norm (f x) \ norm x * B" using lf unfolding linear_conv_bounded_linear by (rule bounded_linear.pos_bounded) with that show ?thesis by (auto simp: mult.commute) qed lemma linear_invertible_bounded_below_pos: fixes f :: "'a::real_normed_vector \ 'b::euclidean_space" assumes "linear f" "linear g" "g \ f = id" obtains B where "B > 0" "\x. B * norm x \ norm(f x)" proof - obtain B where "B > 0" and B: "\x. norm (g x) \ B * norm x" using linear_bounded_pos [OF \linear g\] by blast show thesis proof show "0 < 1/B" by (simp add: \B > 0\) show "1/B * norm x \ norm (f x)" for x proof - have "1/B * norm x = 1/B * norm (g (f x))" using assms by (simp add: pointfree_idE) also have "\ \ norm (f x)" using B [of "f x"] by (simp add: \B > 0\ mult.commute pos_divide_le_eq) finally show ?thesis . qed qed qed lemma linear_inj_bounded_below_pos: fixes f :: "'a::real_normed_vector \ 'b::euclidean_space" assumes "linear f" "inj f" obtains B where "B > 0" "\x. B * norm x \ norm(f x)" using linear_injective_left_inverse [OF assms] linear_invertible_bounded_below_pos assms by blast lemma bounded_linearI': fixes f ::"'a::euclidean_space \ 'b::real_normed_vector" assumes "\x y. f (x + y) = f x + f y" and "\c x. f (c *\<^sub>R x) = c *\<^sub>R f x" shows "bounded_linear f" using assms linearI linear_conv_bounded_linear by blast lemma bilinear_bounded: fixes h :: "'m::euclidean_space \ 'n::euclidean_space \ 'k::real_normed_vector" assumes bh: "bilinear h" shows "\B. \x y. norm (h x y) \ B * norm x * norm y" proof (clarify intro!: exI[of _ "\i\Basis. \j\Basis. norm (h i j)"]) fix x :: 'm fix y :: 'n have "norm (h x y) = norm (h (sum (\i. (x \ i) *\<^sub>R i) Basis) (sum (\i. (y \ i) *\<^sub>R i) Basis))" by (simp add: euclidean_representation) also have "\ = norm (sum (\ (i,j). h ((x \ i) *\<^sub>R i) ((y \ j) *\<^sub>R j)) (Basis \ Basis))" unfolding bilinear_sum[OF bh] .. finally have th: "norm (h x y) = \" . have "\i j. \i \ Basis; j \ Basis\ \ \x \ i\ * (\y \ j\ * norm (h i j)) \ norm x * (norm y * norm (h i j))" by (auto simp add: zero_le_mult_iff Basis_le_norm mult_mono) then show "norm (h x y) \ (\i\Basis. \j\Basis. norm (h i j)) * norm x * norm y" unfolding sum_distrib_right th sum.cartesian_product by (clarsimp simp add: bilinear_rmul[OF bh] bilinear_lmul[OF bh] field_simps simp del: scaleR_scaleR intro!: sum_norm_le) qed lemma bilinear_conv_bounded_bilinear: fixes h :: "'a::euclidean_space \ 'b::euclidean_space \ 'c::real_normed_vector" shows "bilinear h \ bounded_bilinear h" proof assume "bilinear h" show "bounded_bilinear h" proof fix x y z show "h (x + y) z = h x z + h y z" using \bilinear h\ unfolding bilinear_def linear_iff by simp next fix x y z show "h x (y + z) = h x y + h x z" using \bilinear h\ unfolding bilinear_def linear_iff by simp next show "h (scaleR r x) y = scaleR r (h x y)" "h x (scaleR r y) = scaleR r (h x y)" for r x y using \bilinear h\ unfolding bilinear_def linear_iff by simp_all next have "\B. \x y. norm (h x y) \ B * norm x * norm y" using \bilinear h\ by (rule bilinear_bounded) then show "\K. \x y. norm (h x y) \ norm x * norm y * K" by (simp add: ac_simps) qed next assume "bounded_bilinear h" then interpret h: bounded_bilinear h . show "bilinear h" unfolding bilinear_def linear_conv_bounded_linear using h.bounded_linear_left h.bounded_linear_right by simp qed lemma bilinear_bounded_pos: fixes h :: "'a::euclidean_space \ 'b::euclidean_space \ 'c::real_normed_vector" assumes bh: "bilinear h" shows "\B > 0. \x y. norm (h x y) \ B * norm x * norm y" proof - have "\B > 0. \x y. norm (h x y) \ norm x * norm y * B" using bh [unfolded bilinear_conv_bounded_bilinear] by (rule bounded_bilinear.pos_bounded) then show ?thesis by (simp only: ac_simps) qed lemma bounded_linear_imp_has_derivative: "bounded_linear f \ (f has_derivative f) net" by (auto simp add: has_derivative_def linear_diff linear_linear linear_def dest: bounded_linear.linear) lemma linear_imp_has_derivative: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" shows "linear f \ (f has_derivative f) net" by (simp add: bounded_linear_imp_has_derivative linear_conv_bounded_linear) lemma bounded_linear_imp_differentiable: "bounded_linear f \ f differentiable net" using bounded_linear_imp_has_derivative differentiable_def by blast lemma linear_imp_differentiable: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" shows "linear f \ f differentiable net" by (metis linear_imp_has_derivative differentiable_def) lemma of_real_differentiable [simp,derivative_intros]: "of_real differentiable F" by (simp add: bounded_linear_imp_differentiable bounded_linear_of_real) subsection\<^marker>\tag unimportant\ \We continue\ lemma independent_bound: fixes S :: "'a::euclidean_space set" shows "independent S \ finite S \ card S \ DIM('a)" by (metis dim_subset_UNIV finiteI_independent dim_span_eq_card_independent) lemmas independent_imp_finite = finiteI_independent corollary\<^marker>\tag unimportant\ independent_card_le: fixes S :: "'a::euclidean_space set" assumes "independent S" shows "card S \ DIM('a)" using assms independent_bound by auto lemma dependent_biggerset: fixes S :: "'a::euclidean_space set" shows "(finite S \ card S > DIM('a)) \ dependent S" by (metis independent_bound not_less) text \Picking an orthogonal replacement for a spanning set.\ lemma vector_sub_project_orthogonal: fixes b x :: "'a::euclidean_space" shows "b \ (x - ((b \ x) / (b \ b)) *\<^sub>R b) = 0" unfolding inner_simps by auto lemma pairwise_orthogonal_insert: assumes "pairwise orthogonal S" and "\y. y \ S \ orthogonal x y" shows "pairwise orthogonal (insert x S)" using assms unfolding pairwise_def by (auto simp add: orthogonal_commute) lemma basis_orthogonal: fixes B :: "'a::real_inner set" assumes fB: "finite B" shows "\C. finite C \ card C \ card B \ span C = span B \ pairwise orthogonal C" (is " \C. ?P B C") using fB proof (induct rule: finite_induct) case empty then show ?case apply (rule exI[where x="{}"]) apply (auto simp add: pairwise_def) done next case (insert a B) note fB = \finite B\ and aB = \a \ B\ from \\C. finite C \ card C \ card B \ span C = span B \ pairwise orthogonal C\ obtain C where C: "finite C" "card C \ card B" "span C = span B" "pairwise orthogonal C" by blast let ?a = "a - sum (\x. (x \ a / (x \ x)) *\<^sub>R x) C" let ?C = "insert ?a C" from C(1) have fC: "finite ?C" by simp from fB aB C(1,2) have cC: "card ?C \ card (insert a B)" by (simp add: card_insert_if) { fix x k have th0: "\(a::'a) b c. a - (b - c) = c + (a - b)" by (simp add: field_simps) have "x - k *\<^sub>R (a - (\x\C. (x \ a / (x \ x)) *\<^sub>R x)) \ span C \ x - k *\<^sub>R a \ span C" apply (simp only: scaleR_right_diff_distrib th0) apply (rule span_add_eq) apply (rule span_scale) apply (rule span_sum) apply (rule span_scale) apply (rule span_base) apply assumption done } then have SC: "span ?C = span (insert a B)" unfolding set_eq_iff span_breakdown_eq C(3)[symmetric] by auto { fix y assume yC: "y \ C" then have Cy: "C = insert y (C - {y})" by blast have fth: "finite (C - {y})" using C by simp have "orthogonal ?a y" unfolding orthogonal_def unfolding inner_diff inner_sum_left right_minus_eq unfolding sum.remove [OF \finite C\ \y \ C\] apply (clarsimp simp add: inner_commute[of y a]) apply (rule sum.neutral) apply clarsimp apply (rule C(4)[unfolded pairwise_def orthogonal_def, rule_format]) using \y \ C\ by auto } with \pairwise orthogonal C\ have CPO: "pairwise orthogonal ?C" by (rule pairwise_orthogonal_insert) from fC cC SC CPO have "?P (insert a B) ?C" by blast then show ?case by blast qed lemma orthogonal_basis_exists: fixes V :: "('a::euclidean_space) set" shows "\B. independent B \ B \ span V \ V \ span B \ (card B = dim V) \ pairwise orthogonal B" proof - from basis_exists[of V] obtain B where B: "B \ V" "independent B" "V \ span B" "card B = dim V" by force from B have fB: "finite B" "card B = dim V" using independent_bound by auto from basis_orthogonal[OF fB(1)] obtain C where C: "finite C" "card C \ card B" "span C = span B" "pairwise orthogonal C" by blast from C B have CSV: "C \ span V" by (metis span_superset span_mono subset_trans) from span_mono[OF B(3)] C have SVC: "span V \ span C" by (simp add: span_span) from card_le_dim_spanning[OF CSV SVC C(1)] C(2,3) fB have iC: "independent C" by (simp) from C fB have "card C \ dim V" by simp moreover have "dim V \ card C" using span_card_ge_dim[OF CSV SVC C(1)] by simp ultimately have CdV: "card C = dim V" using C(1) by simp from C B CSV CdV iC show ?thesis by auto qed text \Low-dimensional subset is in a hyperplane (weak orthogonal complement).\ lemma span_not_univ_orthogonal: fixes S :: "'a::euclidean_space set" assumes sU: "span S \ UNIV" shows "\a::'a. a \ 0 \ (\x \ span S. a \ x = 0)" proof - from sU obtain a where a: "a \ span S" by blast from orthogonal_basis_exists obtain B where B: "independent B" "B \ span S" "S \ span B" "card B = dim S" "pairwise orthogonal B" by blast from B have fB: "finite B" "card B = dim S" using independent_bound by auto from span_mono[OF B(2)] span_mono[OF B(3)] have sSB: "span S = span B" by (simp add: span_span) let ?a = "a - sum (\b. (a \ b / (b \ b)) *\<^sub>R b) B" have "sum (\b. (a \ b / (b \ b)) *\<^sub>R b) B \ span S" unfolding sSB apply (rule span_sum) apply (rule span_scale) apply (rule span_base) apply assumption done with a have a0:"?a \ 0" by auto have "?a \ x = 0" if "x\span B" for x proof (rule span_induct [OF that]) show "subspace {x. ?a \ x = 0}" by (auto simp add: subspace_def inner_add) next { fix x assume x: "x \ B" from x have B': "B = insert x (B - {x})" by blast have fth: "finite (B - {x})" using fB by simp have "?a \ x = 0" apply (subst B') using fB fth unfolding sum_clauses(2)[OF fth] apply simp unfolding inner_simps apply (clarsimp simp add: inner_add inner_sum_left) apply (rule sum.neutral, rule ballI) apply (simp only: inner_commute) apply (auto simp add: x field_simps intro: B(5)[unfolded pairwise_def orthogonal_def, rule_format]) done } then show "?a \ x = 0" if "x \ B" for x using that by blast qed with a0 show ?thesis unfolding sSB by (auto intro: exI[where x="?a"]) qed lemma span_not_univ_subset_hyperplane: fixes S :: "'a::euclidean_space set" assumes SU: "span S \ UNIV" shows "\ a. a \0 \ span S \ {x. a \ x = 0}" using span_not_univ_orthogonal[OF SU] by auto lemma lowdim_subset_hyperplane: fixes S :: "'a::euclidean_space set" assumes d: "dim S < DIM('a)" shows "\a::'a. a \ 0 \ span S \ {x. a \ x = 0}" proof - { assume "span S = UNIV" then have "dim (span S) = dim (UNIV :: ('a) set)" by simp then have "dim S = DIM('a)" by (metis Euclidean_Space.dim_UNIV dim_span) with d have False by arith } then have th: "span S \ UNIV" by blast from span_not_univ_subset_hyperplane[OF th] show ?thesis . qed lemma linear_eq_stdbasis: fixes f :: "'a::euclidean_space \ _" assumes lf: "linear f" and lg: "linear g" and fg: "\b. b \ Basis \ f b = g b" shows "f = g" using linear_eq_on_span[OF lf lg, of Basis] fg by auto text \Similar results for bilinear functions.\ lemma bilinear_eq: assumes bf: "bilinear f" and bg: "bilinear g" and SB: "S \ span B" and TC: "T \ span C" and "x\S" "y\T" and fg: "\x y. \x \ B; y\ C\ \ f x y = g x y" shows "f x y = g x y" proof - let ?P = "{x. \y\ span C. f x y = g x y}" from bf bg have sp: "subspace ?P" unfolding bilinear_def linear_iff subspace_def bf bg by (auto simp add: span_zero bilinear_lzero[OF bf] bilinear_lzero[OF bg] span_add Ball_def intro: bilinear_ladd[OF bf]) have sfg: "\x. x \ B \ subspace {a. f x a = g x a}" apply (auto simp add: subspace_def) using bf bg unfolding bilinear_def linear_iff apply (auto simp add: span_zero bilinear_rzero[OF bf] bilinear_rzero[OF bg] span_add Ball_def intro: bilinear_ladd[OF bf]) done have "\y\ span C. f x y = g x y" if "x \ span B" for x apply (rule span_induct [OF that sp]) using fg sfg span_induct by blast then show ?thesis using SB TC assms by auto qed lemma bilinear_eq_stdbasis: fixes f :: "'a::euclidean_space \ 'b::euclidean_space \ _" assumes bf: "bilinear f" and bg: "bilinear g" and fg: "\i j. i \ Basis \ j \ Basis \ f i j = g i j" shows "f = g" using bilinear_eq[OF bf bg equalityD2[OF span_Basis] equalityD2[OF span_Basis]] fg by blast subsection \Infinity norm\ definition\<^marker>\tag important\ "infnorm (x::'a::euclidean_space) = Sup {\x \ b\ |b. b \ Basis}" lemma infnorm_set_image: fixes x :: "'a::euclidean_space" shows "{\x \ i\ |i. i \ Basis} = (\i. \x \ i\) ` Basis" by blast lemma infnorm_Max: fixes x :: "'a::euclidean_space" shows "infnorm x = Max ((\i. \x \ i\) ` Basis)" by (simp add: infnorm_def infnorm_set_image cSup_eq_Max) lemma infnorm_set_lemma: fixes x :: "'a::euclidean_space" shows "finite {\x \ i\ |i. i \ Basis}" and "{\x \ i\ |i. i \ Basis} \ {}" unfolding infnorm_set_image by auto lemma infnorm_pos_le: fixes x :: "'a::euclidean_space" shows "0 \ infnorm x" by (simp add: infnorm_Max Max_ge_iff ex_in_conv) lemma infnorm_triangle: fixes x :: "'a::euclidean_space" shows "infnorm (x + y) \ infnorm x + infnorm y" proof - have *: "\a b c d :: real. \a\ \ c \ \b\ \ d \ \a + b\ \ c + d" by simp show ?thesis by (auto simp: infnorm_Max inner_add_left intro!: *) qed lemma infnorm_eq_0: fixes x :: "'a::euclidean_space" shows "infnorm x = 0 \ x = 0" proof - have "infnorm x \ 0 \ x = 0" unfolding infnorm_Max by (simp add: euclidean_all_zero_iff) then show ?thesis using infnorm_pos_le[of x] by simp qed lemma infnorm_0: "infnorm 0 = 0" by (simp add: infnorm_eq_0) lemma infnorm_neg: "infnorm (- x) = infnorm x" unfolding infnorm_def by simp lemma infnorm_sub: "infnorm (x - y) = infnorm (y - x)" by (metis infnorm_neg minus_diff_eq) lemma absdiff_infnorm: "\infnorm x - infnorm y\ \ infnorm (x - y)" proof - have *: "\(nx::real) n ny. nx \ n + ny \ ny \ n + nx \ \nx - ny\ \ n" by arith show ?thesis proof (rule *) from infnorm_triangle[of "x - y" " y"] infnorm_triangle[of "x - y" "-x"] show "infnorm x \ infnorm (x - y) + infnorm y" "infnorm y \ infnorm (x - y) + infnorm x" by (simp_all add: field_simps infnorm_neg) qed qed lemma real_abs_infnorm: "\infnorm x\ = infnorm x" using infnorm_pos_le[of x] by arith lemma Basis_le_infnorm: fixes x :: "'a::euclidean_space" shows "b \ Basis \ \x \ b\ \ infnorm x" by (simp add: infnorm_Max) lemma infnorm_mul: "infnorm (a *\<^sub>R x) = \a\ * infnorm x" unfolding infnorm_Max proof (safe intro!: Max_eqI) let ?B = "(\i. \x \ i\) ` Basis" { fix b :: 'a assume "b \ Basis" then show "\a *\<^sub>R x \ b\ \ \a\ * Max ?B" by (simp add: abs_mult mult_left_mono) next from Max_in[of ?B] obtain b where "b \ Basis" "Max ?B = \x \ b\" by (auto simp del: Max_in) then show "\a\ * Max ((\i. \x \ i\) ` Basis) \ (\i. \a *\<^sub>R x \ i\) ` Basis" by (intro image_eqI[where x=b]) (auto simp: abs_mult) } qed simp lemma infnorm_mul_lemma: "infnorm (a *\<^sub>R x) \ \a\ * infnorm x" unfolding infnorm_mul .. lemma infnorm_pos_lt: "infnorm x > 0 \ x \ 0" using infnorm_pos_le[of x] infnorm_eq_0[of x] by arith text \Prove that it differs only up to a bound from Euclidean norm.\ lemma infnorm_le_norm: "infnorm x \ norm x" by (simp add: Basis_le_norm infnorm_Max) lemma norm_le_infnorm: fixes x :: "'a::euclidean_space" shows "norm x \ sqrt DIM('a) * infnorm x" unfolding norm_eq_sqrt_inner id_def proof (rule real_le_lsqrt[OF inner_ge_zero]) show "sqrt DIM('a) * infnorm x \ 0" by (simp add: zero_le_mult_iff infnorm_pos_le) have "x \ x \ (\b\Basis. x \ b * (x \ b))" by (metis euclidean_inner order_refl) also have "... \ DIM('a) * \infnorm x\\<^sup>2" by (rule sum_bounded_above) (metis Basis_le_infnorm abs_le_square_iff power2_eq_square real_abs_infnorm) also have "... \ (sqrt DIM('a) * infnorm x)\<^sup>2" by (simp add: power_mult_distrib) finally show "x \ x \ (sqrt DIM('a) * infnorm x)\<^sup>2" . qed lemma tendsto_infnorm [tendsto_intros]: assumes "(f \ a) F" shows "((\x. infnorm (f x)) \ infnorm a) F" proof (rule tendsto_compose [OF LIM_I assms]) fix r :: real assume "r > 0" then show "\s>0. \x. x \ a \ norm (x - a) < s \ norm (infnorm x - infnorm a) < r" by (metis real_norm_def le_less_trans absdiff_infnorm infnorm_le_norm) qed text \Equality in Cauchy-Schwarz and triangle inequalities.\ lemma norm_cauchy_schwarz_eq: "x \ y = norm x * norm y \ norm x *\<^sub>R y = norm y *\<^sub>R x" (is "?lhs \ ?rhs") proof (cases "x=0") case True then show ?thesis by auto next case False from inner_eq_zero_iff[of "norm y *\<^sub>R x - norm x *\<^sub>R y"] have "?rhs \ (norm y * (norm y * norm x * norm x - norm x * (x \ y)) - norm x * (norm y * (y \ x) - norm x * norm y * norm y) = 0)" using False unfolding inner_simps by (auto simp add: power2_norm_eq_inner[symmetric] power2_eq_square inner_commute field_simps) also have "\ \ (2 * norm x * norm y * (norm x * norm y - x \ y) = 0)" using False by (simp add: field_simps inner_commute) also have "\ \ ?lhs" using False by auto finally show ?thesis by metis qed lemma norm_cauchy_schwarz_abs_eq: "\x \ y\ = norm x * norm y \ norm x *\<^sub>R y = norm y *\<^sub>R x \ norm x *\<^sub>R y = - norm y *\<^sub>R x" (is "?lhs \ ?rhs") proof - have th: "\(x::real) a. a \ 0 \ \x\ = a \ x = a \ x = - a" by arith have "?rhs \ norm x *\<^sub>R y = norm y *\<^sub>R x \ norm (- x) *\<^sub>R y = norm y *\<^sub>R (- x)" by simp also have "\ \ (x \ y = norm x * norm y \ (- x) \ y = norm x * norm y)" unfolding norm_cauchy_schwarz_eq[symmetric] unfolding norm_minus_cancel norm_scaleR .. also have "\ \ ?lhs" unfolding th[OF mult_nonneg_nonneg, OF norm_ge_zero[of x] norm_ge_zero[of y]] inner_simps by auto finally show ?thesis .. qed lemma norm_triangle_eq: fixes x y :: "'a::real_inner" shows "norm (x + y) = norm x + norm y \ norm x *\<^sub>R y = norm y *\<^sub>R x" proof (cases "x = 0 \ y = 0") case True then show ?thesis by force next case False then have n: "norm x > 0" "norm y > 0" by auto have "norm (x + y) = norm x + norm y \ (norm (x + y))\<^sup>2 = (norm x + norm y)\<^sup>2" by simp also have "\ \ norm x *\<^sub>R y = norm y *\<^sub>R x" unfolding norm_cauchy_schwarz_eq[symmetric] unfolding power2_norm_eq_inner inner_simps by (simp add: power2_norm_eq_inner[symmetric] power2_eq_square inner_commute field_simps) finally show ?thesis . qed lemma dist_triangle_eq: fixes x y z :: "'a::real_inner" shows "dist x z = dist x y + dist y z \ norm (x - y) *\<^sub>R (y - z) = norm (y - z) *\<^sub>R (x - y)" proof - have *: "x - y + (y - z) = x - z" by auto show ?thesis unfolding dist_norm norm_triangle_eq[of "x - y" "y - z", unfolded *] by (auto simp:norm_minus_commute) qed subsection \Collinearity\ definition\<^marker>\tag important\ collinear :: "'a::real_vector set \ bool" where "collinear S \ (\u. \x \ S. \ y \ S. \c. x - y = c *\<^sub>R u)" lemma collinear_alt: "collinear S \ (\u v. \x \ S. \c. x = u + c *\<^sub>R v)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs unfolding collinear_def by (metis Groups.add_ac(2) diff_add_cancel) next assume ?rhs then obtain u v where *: "\x. x \ S \ \c. x = u + c *\<^sub>R v" - by (auto simp: ) + by auto have "\c. x - y = c *\<^sub>R v" if "x \ S" "y \ S" for x y by (metis *[OF \x \ S\] *[OF \y \ S\] scaleR_left.diff add_diff_cancel_left) then show ?lhs using collinear_def by blast qed lemma collinear: fixes S :: "'a::{perfect_space,real_vector} set" shows "collinear S \ (\u. u \ 0 \ (\x \ S. \ y \ S. \c. x - y = c *\<^sub>R u))" proof - have "\v. v \ 0 \ (\x\S. \y\S. \c. x - y = c *\<^sub>R v)" if "\x\S. \y\S. \c. x - y = c *\<^sub>R u" "u=0" for u proof - have "\x\S. \y\S. x = y" using that by auto moreover obtain v::'a where "v \ 0" using UNIV_not_singleton [of 0] by auto ultimately have "\x\S. \y\S. \c. x - y = c *\<^sub>R v" by auto then show ?thesis using \v \ 0\ by blast qed then show ?thesis apply (clarsimp simp: collinear_def) by (metis scaleR_zero_right vector_fraction_eq_iff) qed lemma collinear_subset: "\collinear T; S \ T\ \ collinear S" by (meson collinear_def subsetCE) lemma collinear_empty [iff]: "collinear {}" by (simp add: collinear_def) lemma collinear_sing [iff]: "collinear {x}" by (simp add: collinear_def) lemma collinear_2 [iff]: "collinear {x, y}" apply (simp add: collinear_def) apply (rule exI[where x="x - y"]) by (metis minus_diff_eq scaleR_left.minus scaleR_one) lemma collinear_lemma: "collinear {0, x, y} \ x = 0 \ y = 0 \ (\c. y = c *\<^sub>R x)" (is "?lhs \ ?rhs") proof (cases "x = 0 \ y = 0") case True then show ?thesis by (auto simp: insert_commute) next case False show ?thesis proof assume h: "?lhs" then obtain u where u: "\ x\ {0,x,y}. \y\ {0,x,y}. \c. x - y = c *\<^sub>R u" unfolding collinear_def by blast from u[rule_format, of x 0] u[rule_format, of y 0] obtain cx and cy where cx: "x = cx *\<^sub>R u" and cy: "y = cy *\<^sub>R u" by auto from cx cy False have cx0: "cx \ 0" and cy0: "cy \ 0" by auto let ?d = "cy / cx" from cx cy cx0 have "y = ?d *\<^sub>R x" by simp then show ?rhs using False by blast next assume h: "?rhs" then obtain c where c: "y = c *\<^sub>R x" using False by blast show ?lhs unfolding collinear_def c apply (rule exI[where x=x]) apply auto apply (rule exI[where x="- 1"], simp) apply (rule exI[where x= "-c"], simp) apply (rule exI[where x=1], simp) apply (rule exI[where x="1 - c"], simp add: scaleR_left_diff_distrib) apply (rule exI[where x="c - 1"], simp add: scaleR_left_diff_distrib) done qed qed lemma collinear_iff_Reals: "collinear {0::complex,w,z} \ z/w \ \" proof show "z/w \ \ \ collinear {0,w,z}" by (metis Reals_cases collinear_lemma nonzero_divide_eq_eq scaleR_conv_of_real) qed (auto simp: collinear_lemma scaleR_conv_of_real) lemma collinear_scaleR_iff: "collinear {0, \ *\<^sub>R w, \ *\<^sub>R z} \ collinear {0,w,z} \ \=0 \ \=0" (is "?lhs = ?rhs") proof (cases "\=0 \ \=0") case False then have "(\c. \ *\<^sub>R z = (c * \) *\<^sub>R w) = (\c. z = c *\<^sub>R w)" by (metis mult.commute scaleR_scaleR vector_fraction_eq_iff) then show ?thesis by (auto simp add: collinear_lemma) qed (auto simp: collinear_lemma) lemma norm_cauchy_schwarz_equal: "\x \ y\ = norm x * norm y \ collinear {0, x, y}" proof (cases "x=0") case True then show ?thesis by (auto simp: insert_commute) next case False then have nnz: "norm x \ 0" by auto show ?thesis proof assume "\x \ y\ = norm x * norm y" then show "collinear {0, x, y}" unfolding norm_cauchy_schwarz_abs_eq collinear_lemma by (meson eq_vector_fraction_iff nnz) next assume "collinear {0, x, y}" with False show "\x \ y\ = norm x * norm y" unfolding norm_cauchy_schwarz_abs_eq collinear_lemma by (auto simp: abs_if) qed qed lemma norm_triangle_eq_imp_collinear: fixes x y :: "'a::real_inner" assumes "norm (x + y) = norm x + norm y" shows "collinear{0,x,y}" proof (cases "x = 0 \ y = 0") case False with assms show ?thesis by (meson norm_cauchy_schwarz_abs_eq norm_cauchy_schwarz_equal norm_triangle_eq) qed (use collinear_lemma in blast) subsection\Properties of special hyperplanes\ lemma subspace_hyperplane: "subspace {x. a \ x = 0}" by (simp add: subspace_def inner_right_distrib) lemma subspace_hyperplane2: "subspace {x. x \ a = 0}" by (simp add: inner_commute inner_right_distrib subspace_def) lemma special_hyperplane_span: fixes S :: "'n::euclidean_space set" assumes "k \ Basis" shows "{x. k \ x = 0} = span (Basis - {k})" proof - have *: "x \ span (Basis - {k})" if "k \ x = 0" for x proof - have "x = (\b\Basis. (x \ b) *\<^sub>R b)" by (simp add: euclidean_representation) also have "... = (\b \ Basis - {k}. (x \ b) *\<^sub>R b)" by (auto simp: sum.remove [of _ k] inner_commute assms that) finally have "x = (\b\Basis - {k}. (x \ b) *\<^sub>R b)" . then show ?thesis by (simp add: span_finite) qed show ?thesis apply (rule span_subspace [symmetric]) using assms apply (auto simp: inner_not_same_Basis intro: * subspace_hyperplane) done qed lemma dim_special_hyperplane: fixes k :: "'n::euclidean_space" shows "k \ Basis \ dim {x. k \ x = 0} = DIM('n) - 1" apply (simp add: special_hyperplane_span) apply (rule dim_unique [OF subset_refl]) apply (auto simp: independent_substdbasis) apply (metis member_remove remove_def span_base) done proposition dim_hyperplane: fixes a :: "'a::euclidean_space" assumes "a \ 0" shows "dim {x. a \ x = 0} = DIM('a) - 1" proof - have span0: "span {x. a \ x = 0} = {x. a \ x = 0}" by (rule span_unique) (auto simp: subspace_hyperplane) then obtain B where "independent B" and Bsub: "B \ {x. a \ x = 0}" and subspB: "{x. a \ x = 0} \ span B" and card0: "(card B = dim {x. a \ x = 0})" and ortho: "pairwise orthogonal B" using orthogonal_basis_exists by metis with assms have "a \ span B" by (metis (mono_tags, lifting) span_eq inner_eq_zero_iff mem_Collect_eq span0) then have ind: "independent (insert a B)" by (simp add: \independent B\ independent_insert) have "finite B" using \independent B\ independent_bound by blast have "UNIV \ span (insert a B)" proof fix y::'a obtain r z where z: "y = r *\<^sub>R a + z" "a \ z = 0" apply (rule_tac r="(a \ y) / (a \ a)" and z = "y - ((a \ y) / (a \ a)) *\<^sub>R a" in that) using assms by (auto simp: algebra_simps) show "y \ span (insert a B)" by (metis (mono_tags, lifting) z Bsub span_eq_iff add_diff_cancel_left' mem_Collect_eq span0 span_breakdown_eq span_subspace subspB) qed then have dima: "DIM('a) = dim(insert a B)" by (metis independent_Basis span_Basis dim_eq_card top.extremum_uniqueI) then show ?thesis by (metis (mono_tags, lifting) Bsub Diff_insert_absorb \a \ span B\ ind card0 card_Diff_singleton dim_span indep_card_eq_dim_span insertI1 subsetCE subspB) qed lemma lowdim_eq_hyperplane: fixes S :: "'a::euclidean_space set" assumes "dim S = DIM('a) - 1" obtains a where "a \ 0" and "span S = {x. a \ x = 0}" proof - have dimS: "dim S < DIM('a)" by (simp add: assms) then obtain b where b: "b \ 0" "span S \ {a. b \ a = 0}" using lowdim_subset_hyperplane [of S] by fastforce show ?thesis apply (rule that[OF b(1)]) apply (rule subspace_dim_equal) by (auto simp: assms b dim_hyperplane subspace_hyperplane) qed lemma dim_eq_hyperplane: fixes S :: "'n::euclidean_space set" shows "dim S = DIM('n) - 1 \ (\a. a \ 0 \ span S = {x. a \ x = 0})" by (metis One_nat_def dim_hyperplane dim_span lowdim_eq_hyperplane) subsection\ Orthogonal bases and Gram-Schmidt process\ lemma pairwise_orthogonal_independent: assumes "pairwise orthogonal S" and "0 \ S" shows "independent S" proof - have 0: "\x y. \x \ y; x \ S; y \ S\ \ x \ y = 0" using assms by (simp add: pairwise_def orthogonal_def) have "False" if "a \ S" and a: "a \ span (S - {a})" for a proof - obtain T U where "T \ S - {a}" "a = (\v\T. U v *\<^sub>R v)" using a by (force simp: span_explicit) then have "a \ a = a \ (\v\T. U v *\<^sub>R v)" by simp also have "... = 0" apply (simp add: inner_sum_right) apply (rule comm_monoid_add_class.sum.neutral) by (metis "0" DiffE \T \ S - {a}\ mult_not_zero singletonI subsetCE \a \ S\) finally show ?thesis using \0 \ S\ \a \ S\ by auto qed then show ?thesis by (force simp: dependent_def) qed lemma pairwise_orthogonal_imp_finite: fixes S :: "'a::euclidean_space set" assumes "pairwise orthogonal S" shows "finite S" proof - have "independent (S - {0})" apply (rule pairwise_orthogonal_independent) apply (metis Diff_iff assms pairwise_def) by blast then show ?thesis by (meson independent_imp_finite infinite_remove) qed lemma subspace_orthogonal_to_vector: "subspace {y. orthogonal x y}" by (simp add: subspace_def orthogonal_clauses) lemma subspace_orthogonal_to_vectors: "subspace {y. \x \ S. orthogonal x y}" by (simp add: subspace_def orthogonal_clauses) lemma orthogonal_to_span: assumes a: "a \ span S" and x: "\y. y \ S \ orthogonal x y" shows "orthogonal x a" by (metis a orthogonal_clauses(1,2,4) span_induct_alt x) proposition Gram_Schmidt_step: fixes S :: "'a::euclidean_space set" assumes S: "pairwise orthogonal S" and x: "x \ span S" shows "orthogonal x (a - (\b\S. (b \ a / (b \ b)) *\<^sub>R b))" proof - have "finite S" by (simp add: S pairwise_orthogonal_imp_finite) have "orthogonal (a - (\b\S. (b \ a / (b \ b)) *\<^sub>R b)) x" if "x \ S" for x proof - have "a \ x = (\y\S. if y = x then y \ a else 0)" by (simp add: \finite S\ inner_commute that) also have "... = (\b\S. b \ a * (b \ x) / (b \ b))" apply (rule sum.cong [OF refl], simp) by (meson S orthogonal_def pairwise_def that) finally show ?thesis by (simp add: orthogonal_def algebra_simps inner_sum_left) qed then show ?thesis using orthogonal_to_span orthogonal_commute x by blast qed lemma orthogonal_extension_aux: fixes S :: "'a::euclidean_space set" assumes "finite T" "finite S" "pairwise orthogonal S" shows "\U. pairwise orthogonal (S \ U) \ span (S \ U) = span (S \ T)" using assms proof (induction arbitrary: S) case empty then show ?case by simp (metis sup_bot_right) next case (insert a T) have 0: "\x y. \x \ y; x \ S; y \ S\ \ x \ y = 0" using insert by (simp add: pairwise_def orthogonal_def) define a' where "a' = a - (\b\S. (b \ a / (b \ b)) *\<^sub>R b)" obtain U where orthU: "pairwise orthogonal (S \ insert a' U)" and spanU: "span (insert a' S \ U) = span (insert a' S \ T)" by (rule exE [OF insert.IH [of "insert a' S"]]) (auto simp: Gram_Schmidt_step a'_def insert.prems orthogonal_commute pairwise_orthogonal_insert span_clauses) have orthS: "\x. x \ S \ a' \ x = 0" apply (simp add: a'_def) using Gram_Schmidt_step [OF \pairwise orthogonal S\] apply (force simp: orthogonal_def inner_commute span_superset [THEN subsetD]) done have "span (S \ insert a' U) = span (insert a' (S \ T))" using spanU by simp also have "... = span (insert a (S \ T))" apply (rule eq_span_insert_eq) apply (simp add: a'_def span_neg span_sum span_base span_mul) done also have "... = span (S \ insert a T)" by simp finally show ?case by (rule_tac x="insert a' U" in exI) (use orthU in auto) qed proposition orthogonal_extension: fixes S :: "'a::euclidean_space set" assumes S: "pairwise orthogonal S" obtains U where "pairwise orthogonal (S \ U)" "span (S \ U) = span (S \ T)" proof - obtain B where "finite B" "span B = span T" using basis_subspace_exists [of "span T"] subspace_span by metis with orthogonal_extension_aux [of B S] obtain U where "pairwise orthogonal (S \ U)" "span (S \ U) = span (S \ B)" using assms pairwise_orthogonal_imp_finite by auto with \span B = span T\ show ?thesis by (rule_tac U=U in that) (auto simp: span_Un) qed corollary\<^marker>\tag unimportant\ orthogonal_extension_strong: fixes S :: "'a::euclidean_space set" assumes S: "pairwise orthogonal S" obtains U where "U \ (insert 0 S) = {}" "pairwise orthogonal (S \ U)" "span (S \ U) = span (S \ T)" proof - obtain U where "pairwise orthogonal (S \ U)" "span (S \ U) = span (S \ T)" using orthogonal_extension assms by blast then show ?thesis apply (rule_tac U = "U - (insert 0 S)" in that) apply blast apply (force simp: pairwise_def) apply (metis Un_Diff_cancel Un_insert_left span_redundant span_zero) done qed subsection\Decomposing a vector into parts in orthogonal subspaces\ text\existence of orthonormal basis for a subspace.\ lemma orthogonal_spanningset_subspace: fixes S :: "'a :: euclidean_space set" assumes "subspace S" obtains B where "B \ S" "pairwise orthogonal B" "span B = S" proof - obtain B where "B \ S" "independent B" "S \ span B" "card B = dim S" using basis_exists by blast with orthogonal_extension [of "{}" B] show ?thesis by (metis Un_empty_left assms pairwise_empty span_superset span_subspace that) qed lemma orthogonal_basis_subspace: fixes S :: "'a :: euclidean_space set" assumes "subspace S" obtains B where "0 \ B" "B \ S" "pairwise orthogonal B" "independent B" "card B = dim S" "span B = S" proof - obtain B where "B \ S" "pairwise orthogonal B" "span B = S" using assms orthogonal_spanningset_subspace by blast then show ?thesis apply (rule_tac B = "B - {0}" in that) apply (auto simp: indep_card_eq_dim_span pairwise_subset pairwise_orthogonal_independent elim: pairwise_subset) done qed proposition orthonormal_basis_subspace: fixes S :: "'a :: euclidean_space set" assumes "subspace S" obtains B where "B \ S" "pairwise orthogonal B" and "\x. x \ B \ norm x = 1" and "independent B" "card B = dim S" "span B = S" proof - obtain B where "0 \ B" "B \ S" and orth: "pairwise orthogonal B" and "independent B" "card B = dim S" "span B = S" by (blast intro: orthogonal_basis_subspace [OF assms]) have 1: "(\x. x /\<^sub>R norm x) ` B \ S" using \span B = S\ span_superset span_mul by fastforce have 2: "pairwise orthogonal ((\x. x /\<^sub>R norm x) ` B)" using orth by (force simp: pairwise_def orthogonal_clauses) have 3: "\x. x \ (\x. x /\<^sub>R norm x) ` B \ norm x = 1" by (metis (no_types, lifting) \0 \ B\ image_iff norm_sgn sgn_div_norm) have 4: "independent ((\x. x /\<^sub>R norm x) ` B)" by (metis "2" "3" norm_zero pairwise_orthogonal_independent zero_neq_one) have "inj_on (\x. x /\<^sub>R norm x) B" proof fix x y assume "x \ B" "y \ B" "x /\<^sub>R norm x = y /\<^sub>R norm y" moreover have "\i. i \ B \ norm (i /\<^sub>R norm i) = 1" using 3 by blast ultimately show "x = y" by (metis norm_eq_1 orth orthogonal_clauses(7) orthogonal_commute orthogonal_def pairwise_def zero_neq_one) qed then have 5: "card ((\x. x /\<^sub>R norm x) ` B) = dim S" by (metis \card B = dim S\ card_image) have 6: "span ((\x. x /\<^sub>R norm x) ` B) = S" by (metis "1" "4" "5" assms card_eq_dim independent_imp_finite span_subspace) show ?thesis by (rule that [OF 1 2 3 4 5 6]) qed proposition\<^marker>\tag unimportant\ orthogonal_to_subspace_exists_gen: fixes S :: "'a :: euclidean_space set" assumes "span S \ span T" obtains x where "x \ 0" "x \ span T" "\y. y \ span S \ orthogonal x y" proof - obtain B where "B \ span S" and orthB: "pairwise orthogonal B" and "\x. x \ B \ norm x = 1" and "independent B" "card B = dim S" "span B = span S" by (rule orthonormal_basis_subspace [of "span S", OF subspace_span]) (auto) with assms obtain u where spanBT: "span B \ span T" and "u \ span B" "u \ span T" by auto obtain C where orthBC: "pairwise orthogonal (B \ C)" and spanBC: "span (B \ C) = span (B \ {u})" by (blast intro: orthogonal_extension [OF orthB]) show thesis proof (cases "C \ insert 0 B") case True then have "C \ span B" using span_eq by (metis span_insert_0 subset_trans) moreover have "u \ span (B \ C)" using \span (B \ C) = span (B \ {u})\ span_superset by force ultimately show ?thesis using True \u \ span B\ by (metis Un_insert_left span_insert_0 sup.orderE) next case False then obtain x where "x \ C" "x \ 0" "x \ B" by blast then have "x \ span T" by (metis (no_types, lifting) Un_insert_right Un_upper2 \u \ span T\ spanBT spanBC \u \ span T\ insert_subset span_superset span_mono span_span subsetCE subset_trans sup_bot.comm_neutral) moreover have "orthogonal x y" if "y \ span B" for y using that proof (rule span_induct) show "subspace {a. orthogonal x a}" by (simp add: subspace_orthogonal_to_vector) show "\b. b \ B \ orthogonal x b" by (metis Un_iff \x \ C\ \x \ B\ orthBC pairwise_def) qed ultimately show ?thesis using \x \ 0\ that \span B = span S\ by auto qed qed corollary\<^marker>\tag unimportant\ orthogonal_to_subspace_exists: fixes S :: "'a :: euclidean_space set" assumes "dim S < DIM('a)" obtains x where "x \ 0" "\y. y \ span S \ orthogonal x y" proof - have "span S \ UNIV" by (metis (mono_tags) UNIV_I assms inner_eq_zero_iff less_le lowdim_subset_hyperplane mem_Collect_eq top.extremum_strict top.not_eq_extremum) with orthogonal_to_subspace_exists_gen [of S UNIV] that show ?thesis by (auto) qed corollary\<^marker>\tag unimportant\ orthogonal_to_vector_exists: fixes x :: "'a :: euclidean_space" assumes "2 \ DIM('a)" obtains y where "y \ 0" "orthogonal x y" proof - have "dim {x} < DIM('a)" using assms by auto then show thesis by (rule orthogonal_to_subspace_exists) (simp add: orthogonal_commute span_base that) qed proposition\<^marker>\tag unimportant\ orthogonal_subspace_decomp_exists: fixes S :: "'a :: euclidean_space set" obtains y z where "y \ span S" and "\w. w \ span S \ orthogonal z w" and "x = y + z" proof - obtain T where "0 \ T" "T \ span S" "pairwise orthogonal T" "independent T" "card T = dim (span S)" "span T = span S" using orthogonal_basis_subspace subspace_span by blast let ?a = "\b\T. (b \ x / (b \ b)) *\<^sub>R b" have orth: "orthogonal (x - ?a) w" if "w \ span S" for w by (simp add: Gram_Schmidt_step \pairwise orthogonal T\ \span T = span S\ orthogonal_commute that) show ?thesis apply (rule_tac y = "?a" and z = "x - ?a" in that) apply (meson \T \ span S\ span_scale span_sum subsetCE) apply (fact orth, simp) done qed lemma orthogonal_subspace_decomp_unique: fixes S :: "'a :: euclidean_space set" assumes "x + y = x' + y'" and ST: "x \ span S" "x' \ span S" "y \ span T" "y' \ span T" and orth: "\a b. \a \ S; b \ T\ \ orthogonal a b" shows "x = x' \ y = y'" proof - have "x + y - y' = x'" by (simp add: assms) moreover have "\a b. \a \ span S; b \ span T\ \ orthogonal a b" by (meson orth orthogonal_commute orthogonal_to_span) ultimately have "0 = x' - x" by (metis (full_types) add_diff_cancel_left' ST diff_right_commute orthogonal_clauses(10) orthogonal_clauses(5) orthogonal_self) with assms show ?thesis by auto qed lemma vector_in_orthogonal_spanningset: fixes a :: "'a::euclidean_space" obtains S where "a \ S" "pairwise orthogonal S" "span S = UNIV" by (metis UNIV_I Un_iff empty_iff insert_subset orthogonal_extension pairwise_def pairwise_orthogonal_insert span_UNIV subsetI subset_antisym) lemma vector_in_orthogonal_basis: fixes a :: "'a::euclidean_space" assumes "a \ 0" obtains S where "a \ S" "0 \ S" "pairwise orthogonal S" "independent S" "finite S" "span S = UNIV" "card S = DIM('a)" proof - obtain S where S: "a \ S" "pairwise orthogonal S" "span S = UNIV" using vector_in_orthogonal_spanningset . show thesis proof show "pairwise orthogonal (S - {0})" using pairwise_mono S(2) by blast show "independent (S - {0})" by (simp add: \pairwise orthogonal (S - {0})\ pairwise_orthogonal_independent) show "finite (S - {0})" using \independent (S - {0})\ independent_imp_finite by blast show "card (S - {0}) = DIM('a)" using span_delete_0 [of S] S by (simp add: \independent (S - {0})\ indep_card_eq_dim_span) qed (use S \a \ 0\ in auto) qed lemma vector_in_orthonormal_basis: fixes a :: "'a::euclidean_space" assumes "norm a = 1" obtains S where "a \ S" "pairwise orthogonal S" "\x. x \ S \ norm x = 1" "independent S" "card S = DIM('a)" "span S = UNIV" proof - have "a \ 0" using assms by auto then obtain S where "a \ S" "0 \ S" "finite S" and S: "pairwise orthogonal S" "independent S" "span S = UNIV" "card S = DIM('a)" by (metis vector_in_orthogonal_basis) let ?S = "(\x. x /\<^sub>R norm x) ` S" show thesis proof show "a \ ?S" using \a \ S\ assms image_iff by fastforce next show "pairwise orthogonal ?S" using \pairwise orthogonal S\ by (auto simp: pairwise_def orthogonal_def) show "\x. x \ (\x. x /\<^sub>R norm x) ` S \ norm x = 1" using \0 \ S\ by (auto simp: field_split_simps) then show "independent ?S" by (metis \pairwise orthogonal ((\x. x /\<^sub>R norm x) ` S)\ norm_zero pairwise_orthogonal_independent zero_neq_one) have "inj_on (\x. x /\<^sub>R norm x) S" unfolding inj_on_def by (metis (full_types) S(1) \0 \ S\ inverse_nonzero_iff_nonzero norm_eq_zero orthogonal_scaleR orthogonal_self pairwise_def) then show "card ?S = DIM('a)" by (simp add: card_image S) show "span ?S = UNIV" by (metis (no_types) \0 \ S\ \finite S\ \span S = UNIV\ field_class.field_inverse_zero inverse_inverse_eq less_irrefl span_image_scale zero_less_norm_iff) qed qed proposition dim_orthogonal_sum: fixes A :: "'a::euclidean_space set" assumes "\x y. \x \ A; y \ B\ \ x \ y = 0" shows "dim(A \ B) = dim A + dim B" proof - have 1: "\x y. \x \ span A; y \ B\ \ x \ y = 0" by (erule span_induct [OF _ subspace_hyperplane2]; simp add: assms) have "\x y. \x \ span A; y \ span B\ \ x \ y = 0" using 1 by (simp add: span_induct [OF _ subspace_hyperplane]) then have 0: "\x y. \x \ span A; y \ span B\ \ x \ y = 0" by simp have "dim(A \ B) = dim (span (A \ B))" by (simp) also have "span (A \ B) = ((\(a, b). a + b) ` (span A \ span B))" by (auto simp add: span_Un image_def) also have "dim \ = dim {x + y |x y. x \ span A \ y \ span B}" by (auto intro!: arg_cong [where f=dim]) also have "... = dim {x + y |x y. x \ span A \ y \ span B} + dim(span A \ span B)" by (auto simp: dest: 0) also have "... = dim (span A) + dim (span B)" by (rule dim_sums_Int) (auto) also have "... = dim A + dim B" by (simp) finally show ?thesis . qed lemma dim_subspace_orthogonal_to_vectors: fixes A :: "'a::euclidean_space set" assumes "subspace A" "subspace B" "A \ B" shows "dim {y \ B. \x \ A. orthogonal x y} + dim A = dim B" proof - have "dim (span ({y \ B. \x\A. orthogonal x y} \ A)) = dim (span B)" proof (rule arg_cong [where f=dim, OF subset_antisym]) show "span ({y \ B. \x\A. orthogonal x y} \ A) \ span B" by (simp add: \A \ B\ Collect_restrict span_mono) next have *: "x \ span ({y \ B. \x\A. orthogonal x y} \ A)" if "x \ B" for x proof - obtain y z where "x = y + z" "y \ span A" and orth: "\w. w \ span A \ orthogonal z w" using orthogonal_subspace_decomp_exists [of A x] that by auto have "y \ span B" using \y \ span A\ assms(3) span_mono by blast then have "z \ {a \ B. \x. x \ A \ orthogonal x a}" apply simp using \x = y + z\ assms(1) assms(2) orth orthogonal_commute span_add_eq span_eq_iff that by blast then have z: "z \ span {y \ B. \x\A. orthogonal x y}" by (meson span_superset subset_iff) then show ?thesis apply (auto simp: span_Un image_def \x = y + z\ \y \ span A\) using \y \ span A\ add.commute by blast qed show "span B \ span ({y \ B. \x\A. orthogonal x y} \ A)" by (rule span_minimal) (auto intro: * span_minimal) qed then show ?thesis by (metis (no_types, lifting) dim_orthogonal_sum dim_span mem_Collect_eq orthogonal_commute orthogonal_def) qed subsection\Linear functions are (uniformly) continuous on any set\ subsection\<^marker>\tag unimportant\ \Topological properties of linear functions\ lemma linear_lim_0: assumes "bounded_linear f" shows "(f \ 0) (at (0))" proof - interpret f: bounded_linear f by fact have "(f \ f 0) (at 0)" using tendsto_ident_at by (rule f.tendsto) then show ?thesis unfolding f.zero . qed lemma linear_continuous_at: assumes "bounded_linear f" shows "continuous (at a) f" unfolding continuous_at using assms apply (rule bounded_linear.tendsto) apply (rule tendsto_ident_at) done lemma linear_continuous_within: "bounded_linear f \ continuous (at x within s) f" using continuous_at_imp_continuous_at_within linear_continuous_at by blast lemma linear_continuous_on: "bounded_linear f \ continuous_on s f" using continuous_at_imp_continuous_on[of s f] using linear_continuous_at[of f] by auto lemma Lim_linear: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" and h :: "'b \ 'c::real_normed_vector" assumes "(f \ l) F" "linear h" shows "((\x. h(f x)) \ h l) F" proof - obtain B where B: "B > 0" "\x. norm (h x) \ B * norm x" using linear_bounded_pos [OF \linear h\] by blast show ?thesis unfolding tendsto_iff proof (intro allI impI) show "\\<^sub>F x in F. dist (h (f x)) (h l) < e" if "e > 0" for e proof - have "\\<^sub>F x in F. dist (f x) l < e/B" by (simp add: \0 < B\ assms(1) tendstoD that) then show ?thesis unfolding dist_norm proof (rule eventually_mono) show "norm (h (f x) - h l) < e" if "norm (f x - l) < e / B" for x using that B apply (simp add: field_split_simps) by (metis \linear h\ le_less_trans linear_diff) qed qed qed qed lemma linear_continuous_compose: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" and g :: "'b \ 'c::real_normed_vector" assumes "continuous F f" "linear g" shows "continuous F (\x. g(f x))" using assms unfolding continuous_def by (rule Lim_linear) lemma linear_continuous_on_compose: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" and g :: "'b \ 'c::real_normed_vector" assumes "continuous_on S f" "linear g" shows "continuous_on S (\x. g(f x))" using assms by (simp add: continuous_on_eq_continuous_within linear_continuous_compose) text\Also bilinear functions, in composition form\ lemma bilinear_continuous_compose: fixes h :: "'a::euclidean_space \ 'b::euclidean_space \ 'c::real_normed_vector" assumes "continuous F f" "continuous F g" "bilinear h" shows "continuous F (\x. h (f x) (g x))" using assms bilinear_conv_bounded_bilinear bounded_bilinear.continuous by blast lemma bilinear_continuous_on_compose: fixes h :: "'a::euclidean_space \ 'b::euclidean_space \ 'c::real_normed_vector" and f :: "'d::t2_space \ 'a" assumes "continuous_on S f" "continuous_on S g" "bilinear h" shows "continuous_on S (\x. h (f x) (g x))" using assms by (simp add: continuous_on_eq_continuous_within bilinear_continuous_compose) end diff --git a/src/HOL/Analysis/Lipschitz.thy b/src/HOL/Analysis/Lipschitz.thy --- a/src/HOL/Analysis/Lipschitz.thy +++ b/src/HOL/Analysis/Lipschitz.thy @@ -1,839 +1,839 @@ (* Title: HOL/Analysis/Lipschitz.thy Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr Author: Fabian Immler, TU München *) section \Lipschitz Continuity\ theory Lipschitz imports Derivative begin definition\<^marker>\tag important\ lipschitz_on where "lipschitz_on C U f \ (0 \ C \ (\x \ U. \y\U. dist (f x) (f y) \ C * dist x y))" bundle lipschitz_syntax begin notation\<^marker>\tag important\ lipschitz_on ("_-lipschitz'_on" [1000]) end bundle no_lipschitz_syntax begin no_notation lipschitz_on ("_-lipschitz'_on" [1000]) end unbundle lipschitz_syntax lemma lipschitz_onI: "L-lipschitz_on X f" if "\x y. x \ X \ y \ X \ dist (f x) (f y) \ L * dist x y" "0 \ L" using that by (auto simp: lipschitz_on_def) lemma lipschitz_onD: "dist (f x) (f y) \ L * dist x y" if "L-lipschitz_on X f" "x \ X" "y \ X" using that by (auto simp: lipschitz_on_def) lemma lipschitz_on_nonneg: "0 \ L" if "L-lipschitz_on X f" using that by (auto simp: lipschitz_on_def) lemma lipschitz_on_normD: "norm (f x - f y) \ L * norm (x - y)" if "lipschitz_on L X f" "x \ X" "y \ X" using lipschitz_onD[OF that] by (simp add: dist_norm) lemma lipschitz_on_mono: "L-lipschitz_on D f" if "M-lipschitz_on E f" "D \ E" "M \ L" using that by (force simp: lipschitz_on_def intro: order_trans[OF _ mult_right_mono]) lemmas lipschitz_on_subset = lipschitz_on_mono[OF _ _ order_refl] and lipschitz_on_le = lipschitz_on_mono[OF _ order_refl] lemma lipschitz_on_leI: "L-lipschitz_on X f" if "\x y. x \ X \ y \ X \ x \ y \ dist (f x) (f y) \ L * dist x y" "0 \ L" for f::"'a::{linorder_topology, ordered_real_vector, metric_space} \ 'b::metric_space" proof (rule lipschitz_onI) fix x y assume xy: "x \ X" "y \ X" consider "y \ x" | "x \ y" by (rule le_cases) then show "dist (f x) (f y) \ L * dist x y" proof cases case 1 then have "dist (f y) (f x) \ L * dist y x" by (auto intro!: that xy) then show ?thesis by (simp add: dist_commute) qed (auto intro!: that xy) qed fact lemma lipschitz_on_concat: fixes a b c::real assumes f: "L-lipschitz_on {a .. b} f" assumes g: "L-lipschitz_on {b .. c} g" assumes fg: "f b = g b" shows "lipschitz_on L {a .. c} (\x. if x \ b then f x else g x)" (is "lipschitz_on _ _ ?f") proof (rule lipschitz_on_leI) fix x y assume x: "x \ {a..c}" and y: "y \ {a..c}" and xy: "x \ y" consider "x \ b \ b < y" | "x \ b \ y \ b" by arith then show "dist (?f x) (?f y) \ L * dist x y" proof cases case 1 have "dist (f x) (g y) \ dist (f x) (f b) + dist (g b) (g y)" unfolding fg by (rule dist_triangle) also have "dist (f x) (f b) \ L * dist x b" using 1 x by (auto intro!: lipschitz_onD[OF f]) also have "dist (g b) (g y) \ L * dist b y" using 1 x y by (auto intro!: lipschitz_onD[OF g] lipschitz_onD[OF f]) finally have "dist (f x) (g y) \ L * dist x b + L * dist b y" by simp also have "\ = L * (dist x b + dist b y)" by (simp add: algebra_simps) also have "dist x b + dist b y = dist x y" using 1 x y by (auto simp: dist_real_def abs_real_def) finally show ?thesis using 1 by simp next case 2 with lipschitz_onD[OF f, of x y] lipschitz_onD[OF g, of x y] x y xy show ?thesis by (auto simp: fg) qed qed (rule lipschitz_on_nonneg[OF f]) lemma lipschitz_on_concat_max: fixes a b c::real assumes f: "L-lipschitz_on {a .. b} f" assumes g: "M-lipschitz_on {b .. c} g" assumes fg: "f b = g b" shows "(max L M)-lipschitz_on {a .. c} (\x. if x \ b then f x else g x)" proof - have "lipschitz_on (max L M) {a .. b} f" "lipschitz_on (max L M) {b .. c} g" by (auto intro!: lipschitz_on_mono[OF f order_refl] lipschitz_on_mono[OF g order_refl]) from lipschitz_on_concat[OF this fg] show ?thesis . qed subsubsection \Continuity\ proposition lipschitz_on_uniformly_continuous: assumes "L-lipschitz_on X f" shows "uniformly_continuous_on X f" unfolding uniformly_continuous_on_def proof safe fix e::real assume "0 < e" from assms have l: "(L+1)-lipschitz_on X f" by (rule lipschitz_on_mono) auto show "\d>0. \x\X. \x'\X. dist x' x < d \ dist (f x') (f x) < e" using lipschitz_onD[OF l] lipschitz_on_nonneg[OF assms] \0 < e\ by (force intro!: exI[where x="e/(L + 1)"] simp: field_simps) qed proposition lipschitz_on_continuous_on: "continuous_on X f" if "L-lipschitz_on X f" by (rule uniformly_continuous_imp_continuous[OF lipschitz_on_uniformly_continuous[OF that]]) lemma lipschitz_on_continuous_within: "continuous (at x within X) f" if "L-lipschitz_on X f" "x \ X" using lipschitz_on_continuous_on[OF that(1)] that(2) by (auto simp: continuous_on_eq_continuous_within) subsubsection \Differentiable functions\ proposition bounded_derivative_imp_lipschitz: assumes "\x. x \ X \ (f has_derivative f' x) (at x within X)" assumes convex: "convex X" assumes "\x. x \ X \ onorm (f' x) \ C" "0 \ C" shows "C-lipschitz_on X f" proof (rule lipschitz_onI) show "\x y. x \ X \ y \ X \ dist (f x) (f y) \ C * dist x y" by (auto intro!: assms differentiable_bound[unfolded dist_norm[symmetric], OF convex]) qed fact subsubsection\<^marker>\tag unimportant\ \Structural introduction rules\ named_theorems lipschitz_intros "structural introduction rules for Lipschitz controls" lemma lipschitz_on_compose [lipschitz_intros]: "(D * C)-lipschitz_on U (g o f)" if f: "C-lipschitz_on U f" and g: "D-lipschitz_on (f`U) g" proof (rule lipschitz_onI) show "D* C \ 0" using lipschitz_on_nonneg[OF f] lipschitz_on_nonneg[OF g] by auto fix x y assume H: "x \ U" "y \ U" have "dist (g (f x)) (g (f y)) \ D * dist (f x) (f y)" apply (rule lipschitz_onD[OF g]) using H by auto also have "... \ D * C * dist x y" using mult_left_mono[OF lipschitz_onD(1)[OF f H] lipschitz_on_nonneg[OF g]] by auto finally show "dist ((g \ f) x) ((g \ f) y) \ D * C* dist x y" unfolding comp_def by (auto simp add: mult.commute) qed lemma lipschitz_on_compose2: "(D * C)-lipschitz_on U (\x. g (f x))" if "C-lipschitz_on U f" "D-lipschitz_on (f`U) g" using lipschitz_on_compose[OF that] by (simp add: o_def) lemma lipschitz_on_cong[cong]: "C-lipschitz_on U g \ D-lipschitz_on V f" if "C = D" "U = V" "\x. x \ V \ g x = f x" using that by (auto simp: lipschitz_on_def) lemma lipschitz_on_transform: "C-lipschitz_on U g" if "C-lipschitz_on U f" "\x. x \ U \ g x = f x" using that by simp lemma lipschitz_on_empty_iff[simp]: "C-lipschitz_on {} f \ C \ 0" by (auto simp: lipschitz_on_def) lemma lipschitz_on_insert_iff[simp]: "C-lipschitz_on (insert y X) f \ C-lipschitz_on X f \ (\x \ X. dist (f x) (f y) \ C * dist x y)" by (auto simp: lipschitz_on_def dist_commute) lemma lipschitz_on_singleton [lipschitz_intros]: "C \ 0 \ C-lipschitz_on {x} f" and lipschitz_on_empty [lipschitz_intros]: "C \ 0 \ C-lipschitz_on {} f" by simp_all lemma lipschitz_on_id [lipschitz_intros]: "1-lipschitz_on U (\x. x)" by (auto simp: lipschitz_on_def) lemma lipschitz_on_constant [lipschitz_intros]: "0-lipschitz_on U (\x. c)" by (auto simp: lipschitz_on_def) lemma lipschitz_on_add [lipschitz_intros]: fixes f::"'a::metric_space \'b::real_normed_vector" assumes "C-lipschitz_on U f" "D-lipschitz_on U g" shows "(C+D)-lipschitz_on U (\x. f x + g x)" proof (rule lipschitz_onI) show "C + D \ 0" using lipschitz_on_nonneg[OF assms(1)] lipschitz_on_nonneg[OF assms(2)] by auto fix x y assume H: "x \ U" "y \ U" have "dist (f x + g x) (f y + g y) \ dist (f x) (f y) + dist (g x) (g y)" by (simp add: dist_triangle_add) also have "... \ C * dist x y + D * dist x y" using lipschitz_onD(1)[OF assms(1) H] lipschitz_onD(1)[OF assms(2) H] by auto finally show "dist (f x + g x) (f y + g y) \ (C+D) * dist x y" by (auto simp add: algebra_simps) qed lemma lipschitz_on_cmult [lipschitz_intros]: fixes f::"'a::metric_space \ 'b::real_normed_vector" assumes "C-lipschitz_on U f" shows "(abs(a) * C)-lipschitz_on U (\x. a *\<^sub>R f x)" proof (rule lipschitz_onI) show "abs(a) * C \ 0" using lipschitz_on_nonneg[OF assms(1)] by auto fix x y assume H: "x \ U" "y \ U" have "dist (a *\<^sub>R f x) (a *\<^sub>R f y) = abs(a) * dist (f x) (f y)" by (metis dist_norm norm_scaleR real_vector.scale_right_diff_distrib) also have "... \ abs(a) * C * dist x y" using lipschitz_onD(1)[OF assms(1) H] by (simp add: Groups.mult_ac(1) mult_left_mono) finally show "dist (a *\<^sub>R f x) (a *\<^sub>R f y) \ \a\ * C * dist x y" by auto qed lemma lipschitz_on_cmult_real [lipschitz_intros]: fixes f::"'a::metric_space \ real" assumes "C-lipschitz_on U f" shows "(abs(a) * C)-lipschitz_on U (\x. a * f x)" using lipschitz_on_cmult[OF assms] by auto lemma lipschitz_on_cmult_nonneg [lipschitz_intros]: fixes f::"'a::metric_space \ 'b::real_normed_vector" assumes "C-lipschitz_on U f" "a \ 0" shows "(a * C)-lipschitz_on U (\x. a *\<^sub>R f x)" using lipschitz_on_cmult[OF assms(1), of a] assms(2) by auto lemma lipschitz_on_cmult_real_nonneg [lipschitz_intros]: fixes f::"'a::metric_space \ real" assumes "C-lipschitz_on U f" "a \ 0" shows "(a * C)-lipschitz_on U (\x. a * f x)" using lipschitz_on_cmult_nonneg[OF assms] by auto lemma lipschitz_on_cmult_upper [lipschitz_intros]: fixes f::"'a::metric_space \ 'b::real_normed_vector" assumes "C-lipschitz_on U f" "abs(a) \ D" shows "(D * C)-lipschitz_on U (\x. a *\<^sub>R f x)" apply (rule lipschitz_on_mono[OF lipschitz_on_cmult[OF assms(1), of a], of _ "D * C"]) using assms(2) lipschitz_on_nonneg[OF assms(1)] mult_right_mono by auto lemma lipschitz_on_cmult_real_upper [lipschitz_intros]: fixes f::"'a::metric_space \ real" assumes "C-lipschitz_on U f" "abs(a) \ D" shows "(D * C)-lipschitz_on U (\x. a * f x)" using lipschitz_on_cmult_upper[OF assms] by auto lemma lipschitz_on_minus[lipschitz_intros]: fixes f::"'a::metric_space \'b::real_normed_vector" assumes "C-lipschitz_on U f" shows "C-lipschitz_on U (\x. - f x)" by (metis (mono_tags, lifting) assms dist_minus lipschitz_on_def) lemma lipschitz_on_minus_iff[simp]: "L-lipschitz_on X (\x. - f x) \ L-lipschitz_on X f" "L-lipschitz_on X (- f) \ L-lipschitz_on X f" for f::"'a::metric_space \'b::real_normed_vector" using lipschitz_on_minus[of L X f] lipschitz_on_minus[of L X "-f"] by auto lemma lipschitz_on_diff[lipschitz_intros]: fixes f::"'a::metric_space \'b::real_normed_vector" assumes "C-lipschitz_on U f" "D-lipschitz_on U g" shows "(C + D)-lipschitz_on U (\x. f x - g x)" using lipschitz_on_add[OF assms(1) lipschitz_on_minus[OF assms(2)]] by auto lemma lipschitz_on_closure [lipschitz_intros]: assumes "C-lipschitz_on U f" "continuous_on (closure U) f" shows "C-lipschitz_on (closure U) f" proof (rule lipschitz_onI) show "C \ 0" using lipschitz_on_nonneg[OF assms(1)] by simp fix x y assume "x \ closure U" "y \ closure U" obtain u v::"nat \ 'a" where *: "\n. u n \ U" "u \ x" "\n. v n \ U" "v \ y" using \x \ closure U\ \y \ closure U\ unfolding closure_sequential by blast have a: "(\n. f (u n)) \ f x" using *(1) *(2) \x \ closure U\ \continuous_on (closure U) f\ unfolding comp_def continuous_on_closure_sequentially[of U f] by auto have b: "(\n. f (v n)) \ f y" using *(3) *(4) \y \ closure U\ \continuous_on (closure U) f\ unfolding comp_def continuous_on_closure_sequentially[of U f] by auto have l: "(\n. C * dist (u n) (v n) - dist (f (u n)) (f (v n))) \ C * dist x y - dist (f x) (f y)" by (intro tendsto_intros * a b) have "C * dist (u n) (v n) - dist (f (u n)) (f (v n)) \ 0" for n using lipschitz_onD(1)[OF assms(1) \u n \ U\ \v n \ U\] by simp then have "C * dist x y - dist (f x) (f y) \ 0" using LIMSEQ_le_const[OF l, of 0] by auto then show "dist (f x) (f y) \ C * dist x y" by auto qed lemma lipschitz_on_Pair[lipschitz_intros]: assumes f: "L-lipschitz_on A f" assumes g: "M-lipschitz_on A g" shows "(sqrt (L\<^sup>2 + M\<^sup>2))-lipschitz_on A (\a. (f a, g a))" proof (rule lipschitz_onI, goal_cases) case (1 x y) have "dist (f x, g x) (f y, g y) = sqrt ((dist (f x) (f y))\<^sup>2 + (dist (g x) (g y))\<^sup>2)" by (auto simp add: dist_Pair_Pair real_le_lsqrt) also have "\ \ sqrt ((L * dist x y)\<^sup>2 + (M * dist x y)\<^sup>2)" by (auto intro!: real_sqrt_le_mono add_mono power_mono 1 lipschitz_onD f g) also have "\ \ sqrt (L\<^sup>2 + M\<^sup>2) * dist x y" by (auto simp: power_mult_distrib ring_distribs[symmetric] real_sqrt_mult) finally show ?case . qed simp lemma lipschitz_extend_closure: fixes f::"('a::metric_space) \ ('b::complete_space)" assumes "C-lipschitz_on U f" shows "\g. C-lipschitz_on (closure U) g \ (\x\U. g x = f x)" proof - obtain g where g: "\x. x \ U \ g x = f x" "uniformly_continuous_on (closure U) g" using uniformly_continuous_on_extension_on_closure[OF lipschitz_on_uniformly_continuous[OF assms]] by metis have "C-lipschitz_on (closure U) g" apply (rule lipschitz_on_closure, rule lipschitz_on_transform[OF assms]) using g uniformly_continuous_imp_continuous[OF g(2)] by auto then show ?thesis using g(1) by auto qed lemma (in bounded_linear) lipschitz_boundE: obtains B where "B-lipschitz_on A f" proof - from nonneg_bounded obtain B where B: "B \ 0" "\x. norm (f x) \ B * norm x" by (auto simp: ac_simps) have "B-lipschitz_on A f" by (auto intro!: lipschitz_onI B simp: dist_norm diff[symmetric]) thus ?thesis .. qed subsection \Local Lipschitz continuity\ text \Given a function defined on a real interval, it is Lipschitz-continuous if and only if it is locally so, as proved in the following lemmas. It is useful especially for piecewise-defined functions: if each piece is Lipschitz, then so is the whole function. The same goes for functions defined on geodesic spaces, or more generally on geodesic subsets in a metric space (for instance convex subsets in a real vector space), and this follows readily from the real case, but we will not prove it explicitly. We give several variations around this statement. This is essentially a connectedness argument.\ lemma locally_lipschitz_imp_lipschitz_aux: fixes f::"real \ ('a::metric_space)" assumes "a \ b" "continuous_on {a..b} f" "\x. x \ {a.. \y \ {x<..b}. dist (f y) (f x) \ M * (y-x)" shows "dist (f b) (f a) \ M * (b-a)" proof - define A where "A = {x \ {a..b}. dist (f x) (f a) \ M * (x-a)}" have *: "A = (\x. M * (x-a) - dist (f x) (f a))-`{0..} \ {a..b}" unfolding A_def by auto have "a \ A" unfolding A_def using \a \ b\ by auto then have "A \ {}" by auto moreover have "bdd_above A" unfolding A_def by auto moreover have "closed A" unfolding * by (rule closed_vimage_Int, auto intro!: continuous_intros assms) ultimately have "Sup A \ A" by (rule closed_contains_Sup) have "Sup A = b" proof (rule ccontr) assume "Sup A \ b" define x where "x = Sup A" have I: "dist (f x) (f a) \ M * (x-a)" using \Sup A \ A\ x_def A_def by auto have "x \ {a..Sup A \ A\ \Sup A \ b\ A_def by auto then obtain y where J: "y \ {x<..b}" "dist (f y) (f x) \ M * (y-x)" using assms(3) by blast have "dist (f y) (f a) \ dist (f y) (f x) + dist (f x) (f a)" by (rule dist_triangle) also have "... \ M * (y-x) + M * (x-a)" using I J(2) by auto finally have "dist (f y) (f a) \ M * (y-a)" by (auto simp add: algebra_simps) then have "y \ A" unfolding A_def using \y \ {x<..b}\ \x \ {a.. by auto then have "y \ Sup A" by (rule cSup_upper, auto simp: A_def) then show False using \y \ {x<..b}\ x_def by auto qed then show ?thesis using \Sup A \ A\ A_def by auto qed lemma locally_lipschitz_imp_lipschitz: fixes f::"real \ ('a::metric_space)" assumes "continuous_on {a..b} f" "\x y. x \ {a.. y > x \ \z \ {x<..y}. dist (f z) (f x) \ M * (z-x)" "M \ 0" shows "lipschitz_on M {a..b} f" proof (rule lipschitz_onI[OF _ \M \ 0\]) have *: "dist (f t) (f s) \ M * (t-s)" if "s \ t" "s \ {a..b}" "t \ {a..b}" for s t proof (rule locally_lipschitz_imp_lipschitz_aux, simp add: \s \ t\) show "continuous_on {s..t} f" using continuous_on_subset[OF assms(1)] that by auto fix x assume "x \ {s.. {a..z\{x<..t}. dist (f z) (f x) \ M * (z - x)" using assms(2)[OF \x \ {a.., of t] \x \ {s.. by auto qed fix x y assume "x \ {a..b}" "y \ {a..b}" consider "x \ y" | "y \ x" by linarith then show "dist (f x) (f y) \ M * dist x y" apply (cases) using *[OF _ \x \ {a..b}\ \y \ {a..b}\] *[OF _ \y \ {a..b}\ \x \ {a..b}\] by (auto simp add: dist_commute dist_real_def) qed text \We deduce that if a function is Lipschitz on finitely many closed sets on the real line, then it is Lipschitz on any interval contained in their union. The difficulty in the proof is to show that any point \z\ in this interval (except the maximum) has a point arbitrarily close to it on its right which is contained in a common initial closed set. Otherwise, we show that there is a small interval \(z, T)\ which does not intersect any of the initial closed sets, a contradiction.\ proposition lipschitz_on_closed_Union: assumes "\i. i \ I \ lipschitz_on M (U i) f" "\i. i \ I \ closed (U i)" "finite I" "M \ 0" "{u..(v::real)} \ (\i\I. U i)" shows "lipschitz_on M {u..v} f" proof (rule locally_lipschitz_imp_lipschitz[OF _ _ \M \ 0\]) have *: "continuous_on (U i) f" if "i \ I" for i by (rule lipschitz_on_continuous_on[OF assms(1)[OF \i\ I\]]) have "continuous_on (\i\I. U i) f" apply (rule continuous_on_closed_Union) using \finite I\ * assms(2) by auto then show "continuous_on {u..v} f" using \{u..(v::real)} \ (\i\I. U i)\ continuous_on_subset by auto fix z Z assume z: "z \ {u.. v" by auto define T where "T = min Z v" then have T: "T > z" "T \ v" "T \ u" "T \ Z" using z by auto define A where "A = (\i\ I \ {i. U i \ {z<..T} \ {}}. U i \ {z..T})" have a: "closed A" unfolding A_def apply (rule closed_UN) using \finite I\ \\i. i \ I \ closed (U i)\ by auto have b: "bdd_below A" unfolding A_def using \finite I\ by auto have "\i \ I. T \ U i" using \{u..v} \ (\i\I. U i)\ T by auto then have c: "T \ A" unfolding A_def using T by (auto, fastforce) have "Inf A \ z" apply (rule cInf_greatest, auto) using c unfolding A_def by auto moreover have "Inf A \ z" proof (rule ccontr) assume "\(Inf A \ z)" then obtain w where w: "w > z" "w < Inf A" by (meson dense not_le_imp_less) have "Inf A \ T" using a b c by (simp add: cInf_lower) then have "w \ T" using w by auto then have "w \ {u..v}" using w \z \ {u.. T by auto then obtain j where j: "j \ I" "w \ U j" using \{u..v} \ (\i\I. U i)\ by fastforce then have "w \ U j \ {z..T}" "U j \ {z<..T} \ {}" using j T w \w \ T\ by auto then have "w \ A" unfolding A_def using \j \ I\ by auto then have "Inf A \ w" using a b by (simp add: cInf_lower) then show False using w by auto qed ultimately have "Inf A = z" by simp moreover have "Inf A \ A" apply (rule closed_contains_Inf) using a b c by auto ultimately have "z \ A" by simp then obtain i where i: "i \ I" "U i \ {z<..T} \ {}" "z \ U i" unfolding A_def by auto then obtain t where "t \ U i \ {z<..T}" by blast then have "dist (f t) (f z) \ M * (t - z)" using lipschitz_onD(1)[OF assms(1)[of i], of t z] i dist_real_def by auto then show "\t\{z<..Z}. dist (f t) (f z) \ M * (t - z)" using \T \ Z\ \t \ U i \ {z<..T}\ by auto qed subsection \Local Lipschitz continuity (uniform for a family of functions)\ definition\<^marker>\tag important\ local_lipschitz:: "'a::metric_space set \ 'b::metric_space set \ ('a \ 'b \ 'c::metric_space) \ bool" where "local_lipschitz T X f \ \x \ X. \t \ T. \u>0. \L. \t \ cball t u \ T. L-lipschitz_on (cball x u \ X) (f t)" lemma local_lipschitzI: assumes "\t x. t \ T \ x \ X \ \u>0. \L. \t \ cball t u \ T. L-lipschitz_on (cball x u \ X) (f t)" shows "local_lipschitz T X f" using assms unfolding local_lipschitz_def by auto lemma local_lipschitzE: assumes local_lipschitz: "local_lipschitz T X f" assumes "t \ T" "x \ X" obtains u L where "u > 0" "\s. s \ cball t u \ T \ L-lipschitz_on (cball x u \ X) (f s)" using assms local_lipschitz_def by metis lemma local_lipschitz_continuous_on: assumes local_lipschitz: "local_lipschitz T X f" assumes "t \ T" shows "continuous_on X (f t)" unfolding continuous_on_def proof safe fix x assume "x \ X" from local_lipschitzE[OF local_lipschitz \t \ T\ \x \ X\] obtain u L where "0 < u" and L: "\s. s \ cball t u \ T \ L-lipschitz_on (cball x u \ X) (f s)" by metis have "x \ ball x u" using \0 < u\ by simp from lipschitz_on_continuous_on[OF L] have tendsto: "(f t \ f t x) (at x within cball x u \ X)" using \0 < u\ \x \ X\ \t \ T\ by (auto simp: continuous_on_def) moreover have "\\<^sub>F xa in at x. (xa \ cball x u \ X) = (xa \ X)" using eventually_at_ball[OF \0 < u\, of x UNIV] by eventually_elim auto ultimately show "(f t \ f t x) (at x within X)" by (rule Lim_transform_within_set) qed lemma local_lipschitz_compose1: assumes ll: "local_lipschitz (g ` T) X (\t. f t)" assumes g: "continuous_on T g" shows "local_lipschitz T X (\t. f (g t))" proof (rule local_lipschitzI) fix t x assume "t \ T" "x \ X" then have "g t \ g ` T" by simp from local_lipschitzE[OF assms(1) this \x \ X\] obtain u L where "0 < u" and l: "(\s. s \ cball (g t) u \ g ` T \ L-lipschitz_on (cball x u \ X) (f s))" by auto from g[unfolded continuous_on_eq_continuous_within, rule_format, OF \t \ T\, unfolded continuous_within_eps_delta, rule_format, OF \0 < u\] obtain d where d: "d>0" "\x'. x'\T \ dist x' t < d \ dist (g x') (g t) < u" by (auto) show "\u>0. \L. \t\cball t u \ T. L-lipschitz_on (cball x u \ X) (f (g t))" using d \0 < u\ by (fastforce intro: exI[where x="(min d u)/2"] exI[where x=L] intro!: less_imp_le[OF d(2)] lipschitz_on_subset[OF l] simp: dist_commute) qed context fixes T::"'a::metric_space set" and X f assumes local_lipschitz: "local_lipschitz T X f" begin lemma continuous_on_TimesI: assumes y: "\x. x \ X \ continuous_on T (\t. f t x)" shows "continuous_on (T \ X) (\(t, x). f t x)" unfolding continuous_on_iff proof (safe, simp) fix a b and e::real assume H: "a \ T" "b \ X" "0 < e" hence "0 < e/2" by simp from y[unfolded continuous_on_iff, OF \b \ X\, rule_format, OF \a \ T\ \0 < e/2\] obtain d where d: "d > 0" "\t. t \ T \ dist t a < d \ dist (f t b) (f a b) < e/2" by auto from \a : T\ \b \ X\ obtain u L where u: "0 < u" and L: "\t. t \ cball a u \ T \ L-lipschitz_on (cball b u \ X) (f t)" by (erule local_lipschitzE[OF local_lipschitz]) have "a \ cball a u \ T" by (auto simp: \0 < u\ \a \ T\ less_imp_le) from lipschitz_on_nonneg[OF L[OF \a \ cball _ _ \ _\]] have "0 \ L" . let ?d = "Min {d, u, (e/2/(L + 1))}" show "\d>0. \x\T. \y\X. dist (x, y) (a, b) < d \ dist (f x y) (f a b) < e" proof (rule exI[where x = ?d], safe) show "0 < ?d" using \0 \ L\ \0 < u\ \0 < e\ \0 < d\ by (auto intro!: divide_pos_pos ) fix x y assume "x \ T" "y \ X" assume dist_less: "dist (x, y) (a, b) < ?d" have "dist y b \ dist (x, y) (a, b)" using dist_snd_le[of "(x, y)" "(a, b)"] by auto also note dist_less also { note calculation also have "?d \ u" by simp finally have "dist y b < u" . } have "?d \ e/2/(L + 1)" by simp also have "(L + 1) * \ \ e / 2" using \0 < e\ \L \ 0\ by (auto simp: field_split_simps) finally have le1: "(L + 1) * dist y b < e / 2" using \L \ 0\ by simp have "dist x a \ dist (x, y) (a, b)" using dist_fst_le[of "(x, y)" "(a, b)"] by auto also note dist_less finally have "dist x a < ?d" . also have "?d \ d" by simp finally have "dist x a < d" . note \dist x a < ?d\ also have "?d \ u" by simp finally have "dist x a < u" . then have "x \ cball a u \ T" using \x \ T\ by (auto simp: dist_commute) have "dist (f x y) (f a b) \ dist (f x y) (f x b) + dist (f x b) (f a b)" by (rule dist_triangle) also have "(L + 1)-lipschitz_on (cball b u \ X) (f x)" using L[OF \x \ cball a u \ T\] by (rule lipschitz_on_le) simp then have "dist (f x y) (f x b) \ (L + 1) * dist y b" apply (rule lipschitz_onD) subgoal using \y \ X\ \dist y b < u\ by (simp add: dist_commute) subgoal using \0 < u\ \b \ X\ - by (simp add: ) + by simp done also have "(L + 1) * dist y b \ e / 2" using le1 \0 \ L\ by simp also have "dist (f x b) (f a b) < e / 2" by (rule d; fact) also have "e / 2 + e / 2 = e" by simp finally show "dist (f x y) (f a b) < e" by simp qed qed lemma local_lipschitz_compact_implies_lipschitz: assumes "compact X" "compact T" assumes cont: "\x. x \ X \ continuous_on T (\t. f t x)" obtains L where "\t. t \ T \ L-lipschitz_on X (f t)" proof - { assume *: "\n::nat. \(\t\T. n-lipschitz_on X (f t))" { fix n::nat from *[of n] have "\x y t. t \ T \ x \ X \ y \ X \ dist (f t y) (f t x) > n * dist y x" by (force simp: lipschitz_on_def) } then obtain t and x y::"nat \ 'b" where xy: "\n. x n \ X" "\n. y n \ X" and t: "\n. t n \ T" and d: "\n. dist (f (t n) (y n)) (f (t n) (x n)) > n * dist (y n) (x n)" by metis from xy assms obtain lx rx where lx': "lx \ X" "strict_mono (rx :: nat \ nat)" "(x o rx) \ lx" by (metis compact_def) with xy have "\n. (y o rx) n \ X" by auto with assms obtain ly ry where ly': "ly \ X" "strict_mono (ry :: nat \ nat)" "((y o rx) o ry) \ ly" by (metis compact_def) with t have "\n. ((t o rx) o ry) n \ T" by simp with assms obtain lt rt where lt': "lt \ T" "strict_mono (rt :: nat \ nat)" "(((t o rx) o ry) o rt) \ lt" by (metis compact_def) from lx' ly' have lx: "(x o (rx o ry o rt)) \ lx" (is "?x \ _") and ly: "(y o (rx o ry o rt)) \ ly" (is "?y \ _") and lt: "(t o (rx o ry o rt)) \ lt" (is "?t \ _") subgoal by (simp add: LIMSEQ_subseq_LIMSEQ o_assoc lt'(2)) subgoal by (simp add: LIMSEQ_subseq_LIMSEQ ly'(3) o_assoc lt'(2)) subgoal by (simp add: o_assoc lt'(3)) done hence "(\n. dist (?y n) (?x n)) \ dist ly lx" by (metis tendsto_dist) moreover let ?S = "(\(t, x). f t x) ` (T \ X)" have "eventually (\n::nat. n > 0) sequentially" by (metis eventually_at_top_dense) hence "eventually (\n. norm (dist (?y n) (?x n)) \ norm (\diameter ?S\ / n) * 1) sequentially" proof eventually_elim case (elim n) have "0 < rx (ry (rt n))" using \0 < n\ by (metis dual_order.strict_trans1 lt'(2) lx'(2) ly'(2) seq_suble) have compact: "compact ?S" by (auto intro!: compact_continuous_image continuous_on_subset[OF continuous_on_TimesI] compact_Times \compact X\ \compact T\ cont) have "norm (dist (?y n) (?x n)) = dist (?y n) (?x n)" by simp also from this elim d[of "rx (ry (rt n))"] have "\ < dist (f (?t n) (?y n)) (f (?t n) (?x n)) / rx (ry (rt (n)))" using lx'(2) ly'(2) lt'(2) \0 < rx _\ by (auto simp add: field_split_simps strict_mono_def) also have "\ \ diameter ?S / n" proof (rule frac_le) show "diameter ?S \ 0" using compact compact_imp_bounded diameter_ge_0 by blast show "dist (f (?t n) (?y n)) (f (?t n) (?x n)) \ diameter ((\(t,x). f t x) ` (T \ X))" by (metis (no_types) compact compact_imp_bounded diameter_bounded_bound image_eqI mem_Sigma_iff o_apply split_conv t xy(1) xy(2)) show "real n \ real (rx (ry (rt n)))" by (meson le_trans lt'(2) lx'(2) ly'(2) of_nat_mono strict_mono_imp_increasing) qed (use \n > 0\ in auto) also have "\ \ abs (diameter ?S) / n" by (auto intro!: divide_right_mono) finally show ?case by simp qed with _ have "(\n. dist (?y n) (?x n)) \ 0" by (rule tendsto_0_le) (metis tendsto_divide_0[OF tendsto_const] filterlim_at_top_imp_at_infinity filterlim_real_sequentially) ultimately have "lx = ly" using LIMSEQ_unique by fastforce with assms lx' have "lx \ X" by auto from \lt \ T\ this obtain u L where L: "u > 0" "\t. t \ cball lt u \ T \ L-lipschitz_on (cball lx u \ X) (f t)" by (erule local_lipschitzE[OF local_lipschitz]) hence "L \ 0" by (force intro!: lipschitz_on_nonneg \lt \ T\) from L lt ly lx \lx = ly\ have "eventually (\n. ?t n \ ball lt u) sequentially" "eventually (\n. ?y n \ ball lx u) sequentially" "eventually (\n. ?x n \ ball lx u) sequentially" by (auto simp: dist_commute Lim) moreover have "eventually (\n. n > L) sequentially" by (metis filterlim_at_top_dense filterlim_real_sequentially) ultimately have "eventually (\_. False) sequentially" proof eventually_elim case (elim n) hence "dist (f (?t n) (?y n)) (f (?t n) (?x n)) \ L * dist (?y n) (?x n)" using assms xy t unfolding dist_norm[symmetric] by (intro lipschitz_onD[OF L(2)]) (auto) also have "\ \ n * dist (?y n) (?x n)" using elim by (intro mult_right_mono) auto also have "\ \ rx (ry (rt n)) * dist (?y n) (?x n)" by (intro mult_right_mono[OF _ zero_le_dist]) (meson lt'(2) lx'(2) ly'(2) of_nat_le_iff order_trans seq_suble) also have "\ < dist (f (?t n) (?y n)) (f (?t n) (?x n))" by (auto intro!: d) finally show ?case by simp qed hence False by simp } then obtain L where "\t. t \ T \ L-lipschitz_on X (f t)" by metis thus ?thesis .. qed lemma local_lipschitz_subset: assumes "S \ T" "Y \ X" shows "local_lipschitz S Y f" proof (rule local_lipschitzI) fix t x assume "t \ S" "x \ Y" then have "t \ T" "x \ X" using assms by auto from local_lipschitzE[OF local_lipschitz, OF this] obtain u L where u: "0 < u" and L: "\s. s \ cball t u \ T \ L-lipschitz_on (cball x u \ X) (f s)" by blast show "\u>0. \L. \t\cball t u \ S. L-lipschitz_on (cball x u \ Y) (f t)" using assms by (auto intro: exI[where x=u] exI[where x=L] intro!: u lipschitz_on_subset[OF _ Int_mono[OF order_refl \Y \ X\]] L) qed end lemma local_lipschitz_minus: fixes f::"'a::metric_space \ 'b::metric_space \ 'c::real_normed_vector" shows "local_lipschitz T X (\t x. - f t x) = local_lipschitz T X f" by (auto simp: local_lipschitz_def lipschitz_on_minus) lemma local_lipschitz_PairI: assumes f: "local_lipschitz A B (\a b. f a b)" assumes g: "local_lipschitz A B (\a b. g a b)" shows "local_lipschitz A B (\a b. (f a b, g a b))" proof (rule local_lipschitzI) fix t x assume "t \ A" "x \ B" from local_lipschitzE[OF f this] local_lipschitzE[OF g this] obtain u L v M where "0 < u" "(\s. s \ cball t u \ A \ L-lipschitz_on (cball x u \ B) (f s))" "0 < v" "(\s. s \ cball t v \ A \ M-lipschitz_on (cball x v \ B) (g s))" by metis then show "\u>0. \L. \t\cball t u \ A. L-lipschitz_on (cball x u \ B) (\b. (f t b, g t b))" by (intro exI[where x="min u v"]) (force intro: lipschitz_on_subset intro!: lipschitz_on_Pair) qed lemma local_lipschitz_constI: "local_lipschitz S T (\t x. f t)" by (auto simp: intro!: local_lipschitzI lipschitz_on_constant intro: exI[where x=1]) lemma (in bounded_linear) local_lipschitzI: shows "local_lipschitz A B (\_. f)" proof (rule local_lipschitzI, goal_cases) case (1 t x) from lipschitz_boundE[of "(cball x 1 \ B)"] obtain C where "C-lipschitz_on (cball x 1 \ B) f" by auto then show ?case by (auto intro: exI[where x=1]) qed proposition c1_implies_local_lipschitz: fixes T::"real set" and X::"'a::{banach,heine_borel} set" and f::"real \ 'a \ 'a" assumes f': "\t x. t \ T \ x \ X \ (f t has_derivative blinfun_apply (f' (t, x))) (at x)" assumes cont_f': "continuous_on (T \ X) f'" assumes "open T" assumes "open X" shows "local_lipschitz T X f" proof (rule local_lipschitzI) fix t x assume "t \ T" "x \ X" from open_contains_cball[THEN iffD1, OF \open X\, rule_format, OF \x \ X\] obtain u where u: "u > 0" "cball x u \ X" by auto moreover from open_contains_cball[THEN iffD1, OF \open T\, rule_format, OF \t \ T\] obtain v where v: "v > 0" "cball t v \ T" by auto ultimately have "compact (cball t v \ cball x u)" "cball t v \ cball x u \ T \ X" by (auto intro!: compact_Times) then have "compact (f' ` (cball t v \ cball x u))" by (auto intro!: compact_continuous_image continuous_on_subset[OF cont_f']) then obtain B where B: "B > 0" "\s y. s \ cball t v \ y \ cball x u \ norm (f' (s, y)) \ B" by (auto dest!: compact_imp_bounded simp: bounded_pos) have lipschitz: "B-lipschitz_on (cball x (min u v) \ X) (f s)" if s: "s \ cball t v" for s proof - note s also note \cball t v \ T\ finally have deriv: "\y. y \ cball x u \ (f s has_derivative blinfun_apply (f' (s, y))) (at y within cball x u)" using \_ \ X\ by (auto intro!: has_derivative_at_withinI[OF f']) have "norm (f s y - f s z) \ B * norm (y - z)" if "y \ cball x u" "z \ cball x u" for y z using s that by (intro differentiable_bound[OF convex_cball deriv]) (auto intro!: B simp: norm_blinfun.rep_eq[symmetric]) then show ?thesis using \0 < B\ by (auto intro!: lipschitz_onI simp: dist_norm) qed show "\u>0. \L. \t\cball t u \ T. L-lipschitz_on (cball x u \ X) (f t)" by (force intro: exI[where x="min u v"] exI[where x=B] intro!: lipschitz simp: u v) qed end diff --git a/src/HOL/Analysis/T1_Spaces.thy b/src/HOL/Analysis/T1_Spaces.thy --- a/src/HOL/Analysis/T1_Spaces.thy +++ b/src/HOL/Analysis/T1_Spaces.thy @@ -1,701 +1,701 @@ section\T1 and Hausdorff spaces\ theory T1_Spaces imports Product_Topology begin section\T1 spaces with equivalences to many naturally "nice" properties. \ definition t1_space where "t1_space X \ \x \ topspace X. \y \ topspace X. x\y \ (\U. openin X U \ x \ U \ y \ U)" lemma t1_space_expansive: "\topspace Y = topspace X; \U. openin X U \ openin Y U\ \ t1_space X \ t1_space Y" by (metis t1_space_def) lemma t1_space_alt: "t1_space X \ (\x \ topspace X. \y \ topspace X. x\y \ (\U. closedin X U \ x \ U \ y \ U))" by (metis DiffE DiffI closedin_def openin_closedin_eq t1_space_def) lemma t1_space_empty: "topspace X = {} \ t1_space X" by (simp add: t1_space_def) lemma t1_space_derived_set_of_singleton: "t1_space X \ (\x \ topspace X. X derived_set_of {x} = {})" apply (simp add: t1_space_def derived_set_of_def, safe) apply (metis openin_topspace) by force lemma t1_space_derived_set_of_finite: "t1_space X \ (\S. finite S \ X derived_set_of S = {})" proof (intro iffI allI impI) fix S :: "'a set" assume "finite S" then have fin: "finite ((\x. {x}) ` (topspace X \ S))" by blast assume "t1_space X" then have "X derived_set_of (\x \ topspace X \ S. {x}) = {}" unfolding derived_set_of_Union [OF fin] by (auto simp: t1_space_derived_set_of_singleton) then have "X derived_set_of (topspace X \ S) = {}" by simp then show "X derived_set_of S = {}" by simp qed (auto simp: t1_space_derived_set_of_singleton) lemma t1_space_closedin_singleton: "t1_space X \ (\x \ topspace X. closedin X {x})" apply (rule iffI) apply (simp add: closedin_contains_derived_set t1_space_derived_set_of_singleton) using t1_space_alt by auto lemma closedin_t1_singleton: "\t1_space X; a \ topspace X\ \ closedin X {a}" by (simp add: t1_space_closedin_singleton) lemma t1_space_closedin_finite: "t1_space X \ (\S. finite S \ S \ topspace X \ closedin X S)" apply (rule iffI) apply (simp add: closedin_contains_derived_set t1_space_derived_set_of_finite) by (simp add: t1_space_closedin_singleton) lemma closure_of_singleton: "t1_space X \ X closure_of {a} = (if a \ topspace X then {a} else {})" by (simp add: closure_of_eq t1_space_closedin_singleton closure_of_eq_empty_gen) lemma separated_in_singleton: assumes "t1_space X" shows "separatedin X {a} S \ a \ topspace X \ S \ topspace X \ (a \ X closure_of S)" "separatedin X S {a} \ a \ topspace X \ S \ topspace X \ (a \ X closure_of S)" unfolding separatedin_def using assms closure_of closure_of_singleton by fastforce+ lemma t1_space_openin_delete: "t1_space X \ (\U x. openin X U \ x \ U \ openin X (U - {x}))" apply (rule iffI) apply (meson closedin_t1_singleton in_mono openin_diff openin_subset) by (simp add: closedin_def t1_space_closedin_singleton) lemma t1_space_openin_delete_alt: "t1_space X \ (\U x. openin X U \ openin X (U - {x}))" by (metis Diff_empty Diff_insert0 t1_space_openin_delete) lemma t1_space_singleton_Inter_open: "t1_space X \ (\x \ topspace X. \{U. openin X U \ x \ U} = {x})" (is "?P=?Q") and t1_space_Inter_open_supersets: "t1_space X \ (\S. S \ topspace X \ \{U. openin X U \ S \ U} = S)" (is "?P=?R") proof - have "?R \ ?Q" apply clarify apply (drule_tac x="{x}" in spec, simp) done moreover have "?Q \ ?P" apply (clarsimp simp add: t1_space_def) apply (drule_tac x=x in bspec) apply (simp_all add: set_eq_iff) by (metis (no_types, lifting)) moreover have "?P \ ?R" proof (clarsimp simp add: t1_space_closedin_singleton, rule subset_antisym) fix S assume S: "\x\topspace X. closedin X {x}" "S \ topspace X" then show "\ {U. openin X U \ S \ U} \ S" apply clarsimp by (metis Diff_insert_absorb Set.set_insert closedin_def openin_topspace subset_insert) qed force ultimately show "?P=?Q" "?P=?R" by auto qed lemma t1_space_derived_set_of_infinite_openin: "t1_space X \ (\S. X derived_set_of S = {x \ topspace X. \U. x \ U \ openin X U \ infinite(S \ U)})" (is "_ = ?rhs") proof assume "t1_space X" show ?rhs proof safe fix S x U assume "x \ X derived_set_of S" "x \ U" "openin X U" "finite (S \ U)" with \t1_space X\ show "False" apply (simp add: t1_space_derived_set_of_finite) by (metis IntI empty_iff empty_subsetI inf_commute openin_Int_derived_set_of_subset subset_antisym) next fix S x have eq: "(\y. (y \ x) \ y \ S \ y \ T) \ ~((S \ T) \ {x})" for x S T by blast assume "x \ topspace X" "\U. x \ U \ openin X U \ infinite (S \ U)" then show "x \ X derived_set_of S" apply (clarsimp simp add: derived_set_of_def eq) by (meson finite.emptyI finite.insertI finite_subset) qed (auto simp: in_derived_set_of) qed (auto simp: t1_space_derived_set_of_singleton) lemma finite_t1_space_imp_discrete_topology: "\topspace X = U; finite U; t1_space X\ \ X = discrete_topology U" by (metis discrete_topology_unique_derived_set t1_space_derived_set_of_finite) lemma t1_space_subtopology: "t1_space X \ t1_space(subtopology X U)" by (simp add: derived_set_of_subtopology t1_space_derived_set_of_finite) lemma closedin_derived_set_of_gen: "t1_space X \ closedin X (X derived_set_of S)" apply (clarsimp simp add: in_derived_set_of closedin_contains_derived_set derived_set_of_subset_topspace) by (metis DiffD2 insert_Diff insert_iff t1_space_openin_delete) lemma derived_set_of_derived_set_subset_gen: "t1_space X \ X derived_set_of (X derived_set_of S) \ X derived_set_of S" by (meson closedin_contains_derived_set closedin_derived_set_of_gen) lemma subtopology_eq_discrete_topology_gen_finite: "\t1_space X; finite S\ \ subtopology X S = discrete_topology(topspace X \ S)" by (simp add: subtopology_eq_discrete_topology_gen t1_space_derived_set_of_finite) lemma subtopology_eq_discrete_topology_finite: "\t1_space X; S \ topspace X; finite S\ \ subtopology X S = discrete_topology S" by (simp add: subtopology_eq_discrete_topology_eq t1_space_derived_set_of_finite) lemma t1_space_closed_map_image: "\closed_map X Y f; f ` (topspace X) = topspace Y; t1_space X\ \ t1_space Y" by (metis closed_map_def finite_subset_image t1_space_closedin_finite) lemma homeomorphic_t1_space: "X homeomorphic_space Y \ (t1_space X \ t1_space Y)" apply (clarsimp simp add: homeomorphic_space_def) by (meson homeomorphic_eq_everything_map homeomorphic_maps_map t1_space_closed_map_image) proposition t1_space_product_topology: "t1_space (product_topology X I) \ topspace(product_topology X I) = {} \ (\i \ I. t1_space (X i))" proof (cases "topspace(product_topology X I) = {}") case True then show ?thesis using True t1_space_empty by blast next case False then obtain f where f: "f \ (\\<^sub>E i\I. topspace(X i))" by fastforce have "t1_space (product_topology X I) \ (\i\I. t1_space (X i))" proof (intro iffI ballI) show "t1_space (X i)" if "t1_space (product_topology X I)" and "i \ I" for i proof - have clo: "\h. h \ (\\<^sub>E i\I. topspace (X i)) \ closedin (product_topology X I) {h}" using that by (simp add: t1_space_closedin_singleton) show ?thesis unfolding t1_space_closedin_singleton proof clarify show "closedin (X i) {xi}" if "xi \ topspace (X i)" for xi using clo [of "\j \ I. if i=j then xi else f j"] f that \i \ I\ by (fastforce simp add: closedin_product_topology_singleton) qed qed next next show "t1_space (product_topology X I)" if "\i\I. t1_space (X i)" using that by (simp add: t1_space_closedin_singleton Ball_def PiE_iff closedin_product_topology_singleton) qed then show ?thesis using False by blast qed lemma t1_space_prod_topology: "t1_space(prod_topology X Y) \ topspace(prod_topology X Y) = {} \ t1_space X \ t1_space Y" proof (cases "topspace (prod_topology X Y) = {}") case True then show ?thesis by (auto simp: t1_space_empty) next case False have eq: "{(x,y)} = {x} \ {y}" for x y by simp have "t1_space (prod_topology X Y) \ (t1_space X \ t1_space Y)" using False by (force simp: t1_space_closedin_singleton closedin_prod_Times_iff eq simp del: insert_Times_insert) with False show ?thesis by simp qed subsection\Hausdorff Spaces\ definition Hausdorff_space where "Hausdorff_space X \ \x y. x \ topspace X \ y \ topspace X \ (x \ y) \ (\U V. openin X U \ openin X V \ x \ U \ y \ V \ disjnt U V)" lemma Hausdorff_space_expansive: "\Hausdorff_space X; topspace X = topspace Y; \U. openin X U \ openin Y U\ \ Hausdorff_space Y" by (metis Hausdorff_space_def) lemma Hausdorff_space_topspace_empty: "topspace X = {} \ Hausdorff_space X" by (simp add: Hausdorff_space_def) lemma Hausdorff_imp_t1_space: "Hausdorff_space X \ t1_space X" by (metis Hausdorff_space_def disjnt_iff t1_space_def) lemma closedin_derived_set_of: "Hausdorff_space X \ closedin X (X derived_set_of S)" by (simp add: Hausdorff_imp_t1_space closedin_derived_set_of_gen) lemma t1_or_Hausdorff_space: "t1_space X \ Hausdorff_space X \ t1_space X" using Hausdorff_imp_t1_space by blast lemma Hausdorff_space_sing_Inter_opens: "\Hausdorff_space X; a \ topspace X\ \ \{u. openin X u \ a \ u} = {a}" using Hausdorff_imp_t1_space t1_space_singleton_Inter_open by force lemma Hausdorff_space_subtopology: assumes "Hausdorff_space X" shows "Hausdorff_space(subtopology X S)" proof - have *: "disjnt U V \ disjnt (S \ U) (S \ V)" for U V by (simp add: disjnt_iff) from assms show ?thesis apply (simp add: Hausdorff_space_def openin_subtopology_alt) apply (fast intro: * elim!: all_forward) done qed lemma Hausdorff_space_compact_separation: assumes X: "Hausdorff_space X" and S: "compactin X S" and T: "compactin X T" and "disjnt S T" obtains U V where "openin X U" "openin X V" "S \ U" "T \ V" "disjnt U V" proof (cases "S = {}") case True then show thesis by (metis \compactin X T\ compactin_subset_topspace disjnt_empty1 empty_subsetI openin_empty openin_topspace that) next case False have "\x \ S. \U V. openin X U \ openin X V \ x \ U \ T \ V \ disjnt U V" proof fix a assume "a \ S" then have "a \ T" by (meson assms(4) disjnt_iff) have a: "a \ topspace X" using S \a \ S\ compactin_subset_topspace by blast show "\U V. openin X U \ openin X V \ a \ U \ T \ V \ disjnt U V" proof (cases "T = {}") case True then show ?thesis using a disjnt_empty2 openin_empty by blast next case False have "\x \ topspace X - {a}. \U V. openin X U \ openin X V \ x \ U \ a \ V \ disjnt U V" using X a by (simp add: Hausdorff_space_def) then obtain U V where UV: "\x \ topspace X - {a}. openin X (U x) \ openin X (V x) \ x \ U x \ a \ V x \ disjnt (U x) (V x)" by metis with \a \ T\ compactin_subset_topspace [OF T] have Topen: "\W \ U ` T. openin X W" and Tsub: "T \ \ (U ` T)" - by (auto simp: ) + by auto then obtain \ where \: "finite \" "\ \ U ` T" and "T \ \\" using T unfolding compactin_def by meson then obtain F where F: "finite F" "F \ T" "\ = U ` F" and SUF: "T \ \(U ` F)" and "a \ F" using finite_subset_image [OF \] \a \ T\ by (metis subsetD) have U: "\x. \x \ topspace X; x \ a\ \ openin X (U x)" and V: "\x. \x \ topspace X; x \ a\ \ openin X (V x)" and disj: "\x. \x \ topspace X; x \ a\ \ disjnt (U x) (V x)" using UV by blast+ show ?thesis proof (intro exI conjI) have "F \ {}" using False SUF by blast with \a \ F\ show "openin X (\(V ` F))" using F compactin_subset_topspace [OF T] by (force intro: V) show "openin X (\(U ` F))" using F Topen Tsub by (force intro: U) show "disjnt (\(V ` F)) (\(U ` F))" using disj apply (auto simp: disjnt_def) using \F \ T\ \a \ F\ compactin_subset_topspace [OF T] by blast show "a \ (\(V ` F))" using \F \ T\ T UV \a \ T\ compactin_subset_topspace by blast qed (auto simp: SUF) qed qed then obtain U V where UV: "\x \ S. openin X (U x) \ openin X (V x) \ x \ U x \ T \ V x \ disjnt (U x) (V x)" by metis then have "S \ \ (U ` S)" by auto moreover have "\W \ U ` S. openin X W" using UV by blast ultimately obtain I where I: "S \ \ (U ` I)" "I \ S" "finite I" by (metis S compactin_def finite_subset_image) show thesis proof show "openin X (\(U ` I))" using \I \ S\ UV by blast show "openin X (\ (V ` I))" using False UV \I \ S\ \S \ \ (U ` I)\ \finite I\ by blast show "disjnt (\(U ` I)) (\ (V ` I))" by simp (meson UV \I \ S\ disjnt_subset2 in_mono le_INF_iff order_refl) qed (use UV I in auto) qed lemma Hausdorff_space_compact_sets: "Hausdorff_space X \ (\S T. compactin X S \ compactin X T \ disjnt S T \ (\U V. openin X U \ openin X V \ S \ U \ T \ V \ disjnt U V))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (meson Hausdorff_space_compact_separation) next assume R [rule_format]: ?rhs show ?lhs proof (clarsimp simp add: Hausdorff_space_def) fix x y assume "x \ topspace X" "y \ topspace X" "x \ y" then show "\U. openin X U \ (\V. openin X V \ x \ U \ y \ V \ disjnt U V)" using R [of "{x}" "{y}"] by auto qed qed lemma compactin_imp_closedin: assumes X: "Hausdorff_space X" and S: "compactin X S" shows "closedin X S" proof - have "S \ topspace X" by (simp add: assms compactin_subset_topspace) moreover have "\T. openin X T \ x \ T \ T \ topspace X - S" if "x \ topspace X" "x \ S" for x using Hausdorff_space_compact_separation [OF X _ S, of "{x}"] that apply (simp add: disjnt_def) by (metis Diff_mono Diff_triv openin_subset) ultimately show ?thesis using closedin_def openin_subopen by force qed lemma closedin_Hausdorff_singleton: "\Hausdorff_space X; x \ topspace X\ \ closedin X {x}" by (simp add: Hausdorff_imp_t1_space closedin_t1_singleton) lemma closedin_Hausdorff_sing_eq: "Hausdorff_space X \ closedin X {x} \ x \ topspace X" by (meson closedin_Hausdorff_singleton closedin_subset insert_subset) lemma Hausdorff_space_discrete_topology [simp]: "Hausdorff_space (discrete_topology U)" unfolding Hausdorff_space_def apply safe by (metis discrete_topology_unique_alt disjnt_empty2 disjnt_insert2 insert_iff mk_disjoint_insert topspace_discrete_topology) lemma compactin_Int: "\Hausdorff_space X; compactin X S; compactin X T\ \ compactin X (S \ T)" by (simp add: closed_Int_compactin compactin_imp_closedin) lemma finite_topspace_imp_discrete_topology: "\topspace X = U; finite U; Hausdorff_space X\ \ X = discrete_topology U" using Hausdorff_imp_t1_space finite_t1_space_imp_discrete_topology by blast lemma derived_set_of_finite: "\Hausdorff_space X; finite S\ \ X derived_set_of S = {}" using Hausdorff_imp_t1_space t1_space_derived_set_of_finite by auto lemma derived_set_of_singleton: "Hausdorff_space X \ X derived_set_of {x} = {}" by (simp add: derived_set_of_finite) lemma closedin_Hausdorff_finite: "\Hausdorff_space X; S \ topspace X; finite S\ \ closedin X S" by (simp add: compactin_imp_closedin finite_imp_compactin_eq) lemma open_in_Hausdorff_delete: "\Hausdorff_space X; openin X S\ \ openin X (S - {x})" using Hausdorff_imp_t1_space t1_space_openin_delete_alt by auto lemma closedin_Hausdorff_finite_eq: "\Hausdorff_space X; finite S\ \ closedin X S \ S \ topspace X" by (meson closedin_Hausdorff_finite closedin_def) lemma derived_set_of_infinite_openin: "Hausdorff_space X \ X derived_set_of S = {x \ topspace X. \U. x \ U \ openin X U \ infinite(S \ U)}" using Hausdorff_imp_t1_space t1_space_derived_set_of_infinite_openin by fastforce lemma Hausdorff_space_discrete_compactin: "Hausdorff_space X \ S \ X derived_set_of S = {} \ compactin X S \ S \ topspace X \ finite S" using derived_set_of_finite discrete_compactin_eq_finite by fastforce lemma Hausdorff_space_finite_topspace: "Hausdorff_space X \ X derived_set_of (topspace X) = {} \ compact_space X \ finite(topspace X)" using derived_set_of_finite discrete_compact_space_eq_finite by auto lemma derived_set_of_derived_set_subset: "Hausdorff_space X \ X derived_set_of (X derived_set_of S) \ X derived_set_of S" by (simp add: Hausdorff_imp_t1_space derived_set_of_derived_set_subset_gen) lemma Hausdorff_space_injective_preimage: assumes "Hausdorff_space Y" and cmf: "continuous_map X Y f" and "inj_on f (topspace X)" shows "Hausdorff_space X" unfolding Hausdorff_space_def proof clarify fix x y assume x: "x \ topspace X" and y: "y \ topspace X" and "x \ y" then obtain U V where "openin Y U" "openin Y V" "f x \ U" "f y \ V" "disjnt U V" using assms unfolding Hausdorff_space_def continuous_map_def by (meson inj_onD) show "\U V. openin X U \ openin X V \ x \ U \ y \ V \ disjnt U V" proof (intro exI conjI) show "openin X {x \ topspace X. f x \ U}" using \openin Y U\ cmf continuous_map by fastforce show "openin X {x \ topspace X. f x \ V}" using \openin Y V\ cmf openin_continuous_map_preimage by blast show "disjnt {x \ topspace X. f x \ U} {x \ topspace X. f x \ V}" using \disjnt U V\ by (auto simp add: disjnt_def) qed (use x \f x \ U\ y \f y \ V\ in auto) qed lemma homeomorphic_Hausdorff_space: "X homeomorphic_space Y \ Hausdorff_space X \ Hausdorff_space Y" unfolding homeomorphic_space_def homeomorphic_maps_map by (auto simp: homeomorphic_eq_everything_map Hausdorff_space_injective_preimage) lemma Hausdorff_space_retraction_map_image: "\retraction_map X Y r; Hausdorff_space X\ \ Hausdorff_space Y" unfolding retraction_map_def using Hausdorff_space_subtopology homeomorphic_Hausdorff_space retraction_maps_section_image2 by blast lemma compact_Hausdorff_space_optimal: assumes eq: "topspace Y = topspace X" and XY: "\U. openin X U \ openin Y U" and "Hausdorff_space X" "compact_space Y" shows "Y = X" proof - have "\U. closedin X U \ closedin Y U" using XY using topology_finer_closedin [OF eq] by metis have "openin Y S = openin X S" for S by (metis XY assms(3) assms(4) closedin_compact_space compactin_contractive compactin_imp_closedin eq openin_closedin_eq) then show ?thesis by (simp add: topology_eq) qed lemma continuous_map_imp_closed_graph: assumes f: "continuous_map X Y f" and Y: "Hausdorff_space Y" shows "closedin (prod_topology X Y) ((\x. (x,f x)) ` topspace X)" unfolding closedin_def proof show "(\x. (x, f x)) ` topspace X \ topspace (prod_topology X Y)" using continuous_map_def f by fastforce show "openin (prod_topology X Y) (topspace (prod_topology X Y) - (\x. (x, f x)) ` topspace X)" unfolding openin_prod_topology_alt proof (intro allI impI) show "\U V. openin X U \ openin Y V \ x \ U \ y \ V \ U \ V \ topspace (prod_topology X Y) - (\x. (x, f x)) ` topspace X" if "(x,y) \ topspace (prod_topology X Y) - (\x. (x, f x)) ` topspace X" for x y proof - have "x \ topspace X" "y \ topspace Y" "y \ f x" using that by auto moreover have "f x \ topspace Y" by (meson \x \ topspace X\ continuous_map_def f) ultimately obtain U V where UV: "openin Y U" "openin Y V" "f x \ U" "y \ V" "disjnt U V" using Y Hausdorff_space_def by metis show ?thesis proof (intro exI conjI) show "openin X {x \ topspace X. f x \ U}" using \openin Y U\ f openin_continuous_map_preimage by blast show "{x \ topspace X. f x \ U} \ V \ topspace (prod_topology X Y) - (\x. (x, f x)) ` topspace X" using UV by (auto simp: disjnt_iff dest: openin_subset) qed (use UV \x \ topspace X\ in auto) qed qed qed lemma continuous_imp_closed_map: "\continuous_map X Y f; compact_space X; Hausdorff_space Y\ \ closed_map X Y f" by (meson closed_map_def closedin_compact_space compactin_imp_closedin image_compactin) lemma continuous_imp_quotient_map: "\continuous_map X Y f; compact_space X; Hausdorff_space Y; f ` (topspace X) = topspace Y\ \ quotient_map X Y f" by (simp add: continuous_imp_closed_map continuous_closed_imp_quotient_map) lemma continuous_imp_homeomorphic_map: "\continuous_map X Y f; compact_space X; Hausdorff_space Y; f ` (topspace X) = topspace Y; inj_on f (topspace X)\ \ homeomorphic_map X Y f" by (simp add: continuous_imp_closed_map bijective_closed_imp_homeomorphic_map) lemma continuous_imp_embedding_map: "\continuous_map X Y f; compact_space X; Hausdorff_space Y; inj_on f (topspace X)\ \ embedding_map X Y f" by (simp add: continuous_imp_closed_map injective_closed_imp_embedding_map) lemma continuous_inverse_map: assumes "compact_space X" "Hausdorff_space Y" and cmf: "continuous_map X Y f" and gf: "\x. x \ topspace X \ g(f x) = x" and Sf: "S \ f ` (topspace X)" shows "continuous_map (subtopology Y S) X g" proof (rule continuous_map_from_subtopology_mono [OF _ \S \ f ` (topspace X)\]) show "continuous_map (subtopology Y (f ` (topspace X))) X g" unfolding continuous_map_closedin proof (intro conjI ballI allI impI) fix x assume "x \ topspace (subtopology Y (f ` topspace X))" then show "g x \ topspace X" by (auto simp: gf) next fix C assume C: "closedin X C" show "closedin (subtopology Y (f ` topspace X)) {x \ topspace (subtopology Y (f ` topspace X)). g x \ C}" proof (rule compactin_imp_closedin) show "Hausdorff_space (subtopology Y (f ` topspace X))" using Hausdorff_space_subtopology [OF \Hausdorff_space Y\] by blast have "compactin Y (f ` C)" using C cmf image_compactin closedin_compact_space [OF \compact_space X\] by blast moreover have "{x \ topspace Y. x \ f ` topspace X \ g x \ C} = f ` C" using closedin_subset [OF C] cmf by (auto simp: gf continuous_map_def) ultimately have "compactin Y {x \ topspace Y. x \ f ` topspace X \ g x \ C}" by simp then show "compactin (subtopology Y (f ` topspace X)) {x \ topspace (subtopology Y (f ` topspace X)). g x \ C}" by (auto simp add: compactin_subtopology) qed qed qed lemma closed_map_paired_continuous_map_right: "\continuous_map X Y f; Hausdorff_space Y\ \ closed_map X (prod_topology X Y) (\x. (x,f x))" by (simp add: continuous_map_imp_closed_graph embedding_map_graph embedding_imp_closed_map) lemma closed_map_paired_continuous_map_left: assumes f: "continuous_map X Y f" and Y: "Hausdorff_space Y" shows "closed_map X (prod_topology Y X) (\x. (f x,x))" proof - have eq: "(\x. (f x,x)) = (\(a,b). (b,a)) \ (\x. (x,f x))" by auto show ?thesis unfolding eq proof (rule closed_map_compose) show "closed_map X (prod_topology X Y) (\x. (x, f x))" using Y closed_map_paired_continuous_map_right f by blast show "closed_map (prod_topology X Y) (prod_topology Y X) (\(a, b). (b, a))" by (metis homeomorphic_map_swap homeomorphic_imp_closed_map) qed qed lemma proper_map_paired_continuous_map_right: "\continuous_map X Y f; Hausdorff_space Y\ \ proper_map X (prod_topology X Y) (\x. (x,f x))" using closed_injective_imp_proper_map closed_map_paired_continuous_map_right by (metis (mono_tags, lifting) Pair_inject inj_onI) lemma proper_map_paired_continuous_map_left: "\continuous_map X Y f; Hausdorff_space Y\ \ proper_map X (prod_topology Y X) (\x. (f x,x))" using closed_injective_imp_proper_map closed_map_paired_continuous_map_left by (metis (mono_tags, lifting) Pair_inject inj_onI) lemma Hausdorff_space_prod_topology: "Hausdorff_space(prod_topology X Y) \ topspace(prod_topology X Y) = {} \ Hausdorff_space X \ Hausdorff_space Y" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (rule topological_property_of_prod_component) (auto simp: Hausdorff_space_subtopology homeomorphic_Hausdorff_space) next assume R: ?rhs show ?lhs proof (cases "(topspace X \ topspace Y) = {}") case False with R have ne: "topspace X \ {}" "topspace Y \ {}" and X: "Hausdorff_space X" and Y: "Hausdorff_space Y" by auto show ?thesis unfolding Hausdorff_space_def proof clarify fix x y x' y' assume xy: "(x, y) \ topspace (prod_topology X Y)" and xy': "(x',y') \ topspace (prod_topology X Y)" and *: "\U V. openin (prod_topology X Y) U \ openin (prod_topology X Y) V \ (x, y) \ U \ (x', y') \ V \ disjnt U V" have False if "x \ x' \ y \ y'" using that proof assume "x \ x'" then obtain U V where "openin X U" "openin X V" "x \ U" "x' \ V" "disjnt U V" by (metis Hausdorff_space_def X mem_Sigma_iff topspace_prod_topology xy xy') let ?U = "U \ topspace Y" let ?V = "V \ topspace Y" have "openin (prod_topology X Y) ?U" "openin (prod_topology X Y) ?V" by (simp_all add: openin_prod_Times_iff \openin X U\ \openin X V\) moreover have "disjnt ?U ?V" by (simp add: \disjnt U V\) ultimately show False using * \x \ U\ \x' \ V\ xy xy' by (metis SigmaD2 SigmaI topspace_prod_topology) next assume "y \ y'" then obtain U V where "openin Y U" "openin Y V" "y \ U" "y' \ V" "disjnt U V" by (metis Hausdorff_space_def Y mem_Sigma_iff topspace_prod_topology xy xy') let ?U = "topspace X \ U" let ?V = "topspace X \ V" have "openin (prod_topology X Y) ?U" "openin (prod_topology X Y) ?V" by (simp_all add: openin_prod_Times_iff \openin Y U\ \openin Y V\) moreover have "disjnt ?U ?V" by (simp add: \disjnt U V\) ultimately show False using "*" \y \ U\ \y' \ V\ xy xy' by (metis SigmaD1 SigmaI topspace_prod_topology) qed then show "x = x' \ y = y'" by blast qed qed (simp add: Hausdorff_space_topspace_empty) qed lemma Hausdorff_space_product_topology: "Hausdorff_space (product_topology X I) \ (\\<^sub>E i\I. topspace(X i)) = {} \ (\i \ I. Hausdorff_space (X i))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs apply (rule topological_property_of_product_component) apply (blast dest: Hausdorff_space_subtopology homeomorphic_Hausdorff_space)+ done next assume R: ?rhs show ?lhs proof (cases "(\\<^sub>E i\I. topspace(X i)) = {}") case True then show ?thesis by (simp add: Hausdorff_space_topspace_empty) next case False have "\U V. openin (product_topology X I) U \ openin (product_topology X I) V \ f \ U \ g \ V \ disjnt U V" if f: "f \ (\\<^sub>E i\I. topspace (X i))" and g: "g \ (\\<^sub>E i\I. topspace (X i))" and "f \ g" for f g :: "'a \ 'b" proof - obtain m where "f m \ g m" using \f \ g\ by blast then have "m \ I" using f g by fastforce then have "Hausdorff_space (X m)" using False that R by blast then obtain U V where U: "openin (X m) U" and V: "openin (X m) V" and "f m \ U" "g m \ V" "disjnt U V" by (metis Hausdorff_space_def PiE_mem \f m \ g m\ \m \ I\ f g) show ?thesis proof (intro exI conjI) let ?U = "(\\<^sub>E i\I. topspace(X i)) \ {x. x m \ U}" let ?V = "(\\<^sub>E i\I. topspace(X i)) \ {x. x m \ V}" show "openin (product_topology X I) ?U" "openin (product_topology X I) ?V" using \m \ I\ U V by (force simp add: openin_product_topology intro: arbitrary_union_of_inc relative_to_inc finite_intersection_of_inc)+ show "f \ ?U" using \f m \ U\ f by blast show "g \ ?V" using \g m \ V\ g by blast show "disjnt ?U ?V" using \disjnt U V\ by (auto simp: PiE_def Pi_def disjnt_def) qed qed then show ?thesis by (simp add: Hausdorff_space_def) qed qed end diff --git a/src/HOL/Complex_Analysis/Winding_Numbers.thy b/src/HOL/Complex_Analysis/Winding_Numbers.thy --- a/src/HOL/Complex_Analysis/Winding_Numbers.thy +++ b/src/HOL/Complex_Analysis/Winding_Numbers.thy @@ -1,2220 +1,2220 @@ section \Winding numbers\ theory Winding_Numbers imports Cauchy_Integral_Theorem begin subsection \Definition\ definition\<^marker>\tag important\ winding_number_prop :: "[real \ complex, complex, real, real \ complex, complex] \ bool" where "winding_number_prop \ z e p n \ valid_path p \ z \ path_image p \ pathstart p = pathstart \ \ pathfinish p = pathfinish \ \ (\t \ {0..1}. norm(\ t - p t) < e) \ contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" definition\<^marker>\tag important\ winding_number:: "[real \ complex, complex] \ complex" where "winding_number \ z \ SOME n. \e > 0. \p. winding_number_prop \ z e p n" lemma winding_number: assumes "path \" "z \ path_image \" "0 < e" shows "\p. winding_number_prop \ z e p (winding_number \ z)" proof - have "path_image \ \ UNIV - {z}" using assms by blast then obtain d where d: "d>0" and pi_eq: "\h1 h2. valid_path h1 \ valid_path h2 \ (\t\{0..1}. cmod (h1 t - \ t) < d \ cmod (h2 t - \ t) < d) \ pathstart h2 = pathstart h1 \ pathfinish h2 = pathfinish h1 \ path_image h1 \ UNIV - {z} \ path_image h2 \ UNIV - {z} \ (\f. f holomorphic_on UNIV - {z} \ contour_integral h2 f = contour_integral h1 f)" using contour_integral_nearby_ends [of "UNIV - {z}" \] assms by (auto simp: open_delete) then obtain h where h: "polynomial_function h \ pathstart h = pathstart \ \ pathfinish h = pathfinish \ \ (\t \ {0..1}. norm(h t - \ t) < d/2)" using path_approx_polynomial_function [OF \path \\, of "d/2"] d by (metis half_gt_zero_iff) define nn where "nn = 1/(2* pi*\) * contour_integral h (\w. 1/(w - z))" have "\n. \e > 0. \p. winding_number_prop \ z e p n" proof (rule_tac x=nn in exI, clarify) fix e::real assume e: "e>0" obtain p where p: "polynomial_function p \ pathstart p = pathstart \ \ pathfinish p = pathfinish \ \ (\t\{0..1}. cmod (p t - \ t) < min e (d/2))" using path_approx_polynomial_function [OF \path \\, of "min e (d/2)"] d \0 by (metis min_less_iff_conj zero_less_divide_iff zero_less_numeral) have "(\w. 1 / (w - z)) holomorphic_on UNIV - {z}" by (auto simp: intro!: holomorphic_intros) then have "winding_number_prop \ z e p nn" using pi_eq [of h p] h p d by (auto simp: valid_path_polynomial_function norm_minus_commute nn_def winding_number_prop_def) then show "\p. winding_number_prop \ z e p nn" by metis qed then show ?thesis unfolding winding_number_def by (rule someI2_ex) (blast intro: \0) qed lemma winding_number_unique: assumes \: "path \" "z \ path_image \" and pi: "\e. e>0 \ \p. winding_number_prop \ z e p n" shows "winding_number \ z = n" proof - have "path_image \ \ UNIV - {z}" using assms by blast then obtain e where e: "e>0" and pi_eq: "\h1 h2 f. \valid_path h1; valid_path h2; (\t\{0..1}. cmod (h1 t - \ t) < e \ cmod (h2 t - \ t) < e); pathstart h2 = pathstart h1; pathfinish h2 = pathfinish h1; f holomorphic_on UNIV - {z}\ \ contour_integral h2 f = contour_integral h1 f" using contour_integral_nearby_ends [of "UNIV - {z}" \] assms by (auto simp: open_delete) obtain p where p: "winding_number_prop \ z e p n" using pi [OF e] by blast obtain q where q: "winding_number_prop \ z e q (winding_number \ z)" using winding_number [OF \ e] by blast have "2 * complex_of_real pi * \ * n = contour_integral p (\w. 1 / (w - z))" using p by (auto simp: winding_number_prop_def) also have "\ = contour_integral q (\w. 1 / (w - z))" proof (rule pi_eq) show "(\w. 1 / (w - z)) holomorphic_on UNIV - {z}" by (auto intro!: holomorphic_intros) qed (use p q in \auto simp: winding_number_prop_def norm_minus_commute\) also have "\ = 2 * complex_of_real pi * \ * winding_number \ z" using q by (auto simp: winding_number_prop_def) finally have "2 * complex_of_real pi * \ * n = 2 * complex_of_real pi * \ * winding_number \ z" . then show ?thesis by simp qed (*NB not winding_number_prop here due to the loop in p*) lemma winding_number_unique_loop: assumes \: "path \" "z \ path_image \" and loop: "pathfinish \ = pathstart \" and pi: "\e. e>0 \ \p. valid_path p \ z \ path_image p \ pathfinish p = pathstart p \ (\t \ {0..1}. norm (\ t - p t) < e) \ contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" shows "winding_number \ z = n" proof - have "path_image \ \ UNIV - {z}" using assms by blast then obtain e where e: "e>0" and pi_eq: "\h1 h2 f. \valid_path h1; valid_path h2; (\t\{0..1}. cmod (h1 t - \ t) < e \ cmod (h2 t - \ t) < e); pathfinish h1 = pathstart h1; pathfinish h2 = pathstart h2; f holomorphic_on UNIV - {z}\ \ contour_integral h2 f = contour_integral h1 f" using contour_integral_nearby_loops [of "UNIV - {z}" \] assms by (auto simp: open_delete) obtain p where p: "valid_path p \ z \ path_image p \ pathfinish p = pathstart p \ (\t \ {0..1}. norm (\ t - p t) < e) \ contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" using pi [OF e] by blast obtain q where q: "winding_number_prop \ z e q (winding_number \ z)" using winding_number [OF \ e] by blast have "2 * complex_of_real pi * \ * n = contour_integral p (\w. 1 / (w - z))" using p by auto also have "\ = contour_integral q (\w. 1 / (w - z))" proof (rule pi_eq) show "(\w. 1 / (w - z)) holomorphic_on UNIV - {z}" by (auto intro!: holomorphic_intros) qed (use p q loop in \auto simp: winding_number_prop_def norm_minus_commute\) also have "\ = 2 * complex_of_real pi * \ * winding_number \ z" using q by (auto simp: winding_number_prop_def) finally have "2 * complex_of_real pi * \ * n = 2 * complex_of_real pi * \ * winding_number \ z" . then show ?thesis by simp qed proposition winding_number_valid_path: assumes "valid_path \" "z \ path_image \" shows "winding_number \ z = 1/(2*pi*\) * contour_integral \ (\w. 1/(w - z))" by (rule winding_number_unique) (use assms in \auto simp: valid_path_imp_path winding_number_prop_def\) proposition has_contour_integral_winding_number: assumes \: "valid_path \" "z \ path_image \" shows "((\w. 1/(w - z)) has_contour_integral (2*pi*\*winding_number \ z)) \" by (simp add: winding_number_valid_path has_contour_integral_integral contour_integrable_inversediff assms) lemma winding_number_trivial [simp]: "z \ a \ winding_number(linepath a a) z = 0" by (simp add: winding_number_valid_path) lemma winding_number_subpath_trivial [simp]: "z \ g x \ winding_number (subpath x x g) z = 0" by (simp add: path_image_subpath winding_number_valid_path) lemma winding_number_join: assumes \1: "path \1" "z \ path_image \1" and \2: "path \2" "z \ path_image \2" and "pathfinish \1 = pathstart \2" shows "winding_number(\1 +++ \2) z = winding_number \1 z + winding_number \2 z" proof (rule winding_number_unique) show "\p. winding_number_prop (\1 +++ \2) z e p (winding_number \1 z + winding_number \2 z)" if "e > 0" for e proof - obtain p1 where "winding_number_prop \1 z e p1 (winding_number \1 z)" using \0 < e\ \1 winding_number by blast moreover obtain p2 where "winding_number_prop \2 z e p2 (winding_number \2 z)" using \0 < e\ \2 winding_number by blast ultimately have "winding_number_prop (\1+++\2) z e (p1+++p2) (winding_number \1 z + winding_number \2 z)" using assms apply (simp add: winding_number_prop_def not_in_path_image_join contour_integrable_inversediff algebra_simps) apply (auto simp: joinpaths_def) done then show ?thesis by blast qed qed (use assms in \auto simp: not_in_path_image_join\) lemma winding_number_reversepath: assumes "path \" "z \ path_image \" shows "winding_number(reversepath \) z = - (winding_number \ z)" proof (rule winding_number_unique) show "\p. winding_number_prop (reversepath \) z e p (- winding_number \ z)" if "e > 0" for e proof - obtain p where "winding_number_prop \ z e p (winding_number \ z)" using \0 < e\ assms winding_number by blast then have "winding_number_prop (reversepath \) z e (reversepath p) (- winding_number \ z)" using assms unfolding winding_number_prop_def apply (simp add: contour_integral_reversepath contour_integrable_inversediff valid_path_imp_reverse) apply (auto simp: reversepath_def) done then show ?thesis by blast qed qed (use assms in auto) lemma winding_number_shiftpath: assumes \: "path \" "z \ path_image \" and "pathfinish \ = pathstart \" "a \ {0..1}" shows "winding_number(shiftpath a \) z = winding_number \ z" proof (rule winding_number_unique_loop) show "\p. valid_path p \ z \ path_image p \ pathfinish p = pathstart p \ (\t\{0..1}. cmod (shiftpath a \ t - p t) < e) \ contour_integral p (\w. 1 / (w - z)) = 2 * pi * \ * winding_number \ z" if "e > 0" for e proof - obtain p where "winding_number_prop \ z e p (winding_number \ z)" using \0 < e\ assms winding_number by blast then show ?thesis apply (rule_tac x="shiftpath a p" in exI) using assms that apply (auto simp: winding_number_prop_def path_image_shiftpath pathfinish_shiftpath pathstart_shiftpath contour_integral_shiftpath) apply (simp add: shiftpath_def) done qed qed (use assms in \auto simp: path_shiftpath path_image_shiftpath pathfinish_shiftpath pathstart_shiftpath\) lemma winding_number_split_linepath: assumes "c \ closed_segment a b" "z \ closed_segment a b" shows "winding_number(linepath a b) z = winding_number(linepath a c) z + winding_number(linepath c b) z" proof - have "z \ closed_segment a c" "z \ closed_segment c b" using assms by (meson convex_contains_segment convex_segment ends_in_segment subsetCE)+ then show ?thesis using assms by (simp add: winding_number_valid_path contour_integral_split_linepath [symmetric] continuous_on_inversediff field_simps) qed lemma winding_number_cong: "(\t. \0 \ t; t \ 1\ \ p t = q t) \ winding_number p z = winding_number q z" by (simp add: winding_number_def winding_number_prop_def pathstart_def pathfinish_def) lemma winding_number_constI: assumes "c\z" and g: "\t. \0\t; t\1\ \ g t = c" shows "winding_number g z = 0" proof - have "winding_number g z = winding_number (linepath c c) z" using g winding_number_cong by fastforce moreover have "winding_number (linepath c c) z = 0" using \c\z\ by auto ultimately show ?thesis by auto qed lemma winding_number_offset: "winding_number p z = winding_number (\w. p w - z) 0" unfolding winding_number_def proof (intro ext arg_cong [where f = Eps] arg_cong [where f = All] imp_cong refl, safe) fix n e g assume "0 < e" and g: "winding_number_prop p z e g n" then show "\r. winding_number_prop (\w. p w - z) 0 e r n" by (rule_tac x="\t. g t - z" in exI) (force simp: winding_number_prop_def contour_integral_integral valid_path_def path_defs vector_derivative_def has_vector_derivative_diff_const piecewise_C1_differentiable_diff C1_differentiable_imp_piecewise) next fix n e g assume "0 < e" and g: "winding_number_prop (\w. p w - z) 0 e g n" then have "winding_number_prop p z e (\t. g t + z) n" apply (simp add: winding_number_prop_def contour_integral_integral valid_path_def path_defs piecewise_C1_differentiable_add vector_derivative_def has_vector_derivative_add_const C1_differentiable_imp_piecewise) apply (force simp: algebra_simps) done then show "\r. winding_number_prop p z e r n" by metis qed lemma winding_number_negatepath: assumes \: "valid_path \" and 0: "0 \ path_image \" shows "winding_number(uminus \ \) 0 = winding_number \ 0" proof - have "(/) 1 contour_integrable_on \" using "0" \ contour_integrable_inversediff by fastforce then have "((\z. 1/z) has_contour_integral contour_integral \ ((/) 1)) \" by (rule has_contour_integral_integral) then have "((\z. 1 / - z) has_contour_integral - contour_integral \ ((/) 1)) \" using has_contour_integral_neg by auto then have "contour_integral (uminus \ \) ((/) 1) = contour_integral \ ((/) 1)" using \ by (simp add: contour_integral_unique has_contour_integral_negatepath) then show ?thesis using assms by (simp add: winding_number_valid_path valid_path_negatepath image_def path_defs) qed text \A combined theorem deducing several things piecewise.\ lemma winding_number_join_pos_combined: "\valid_path \1; z \ path_image \1; 0 < Re(winding_number \1 z); valid_path \2; z \ path_image \2; 0 < Re(winding_number \2 z); pathfinish \1 = pathstart \2\ \ valid_path(\1 +++ \2) \ z \ path_image(\1 +++ \2) \ 0 < Re(winding_number(\1 +++ \2) z)" by (simp add: valid_path_join path_image_join winding_number_join valid_path_imp_path) subsubsection\<^marker>\tag unimportant\ \Useful sufficient conditions for the winding number to be positive\ lemma Re_winding_number: "\valid_path \; z \ path_image \\ \ Re(winding_number \ z) = Im(contour_integral \ (\w. 1/(w - z))) / (2*pi)" by (simp add: winding_number_valid_path field_simps Re_divide power2_eq_square) lemma winding_number_pos_le: assumes \: "valid_path \" "z \ path_image \" and ge: "\x. \0 < x; x < 1\ \ 0 \ Im (vector_derivative \ (at x) * cnj(\ x - z))" shows "0 \ Re(winding_number \ z)" proof - have ge0: "0 \ Im (vector_derivative \ (at x) / (\ x - z))" if x: "0 < x" "x < 1" for x using ge by (simp add: Complex.Im_divide algebra_simps x) let ?vd = "\x. 1 / (\ x - z) * vector_derivative \ (at x)" let ?int = "\z. contour_integral \ (\w. 1 / (w - z))" have "0 \ Im (?int z)" proof (rule has_integral_component_nonneg [of \, simplified]) show "\x. x \ cbox 0 1 \ 0 \ Im (if 0 < x \ x < 1 then ?vd x else 0)" by (force simp: ge0) have "((\a. 1 / (a - z)) has_contour_integral contour_integral \ (\w. 1 / (w - z))) \" using \ by (simp flip: add: contour_integrable_inversediff has_contour_integral_integral) then have hi: "(?vd has_integral ?int z) (cbox 0 1)" using has_contour_integral by auto show "((\x. if 0 < x \ x < 1 then ?vd x else 0) has_integral ?int z) (cbox 0 1)" by (rule has_integral_spike_interior [OF hi]) simp qed then show ?thesis by (simp add: Re_winding_number [OF \] field_simps) qed lemma winding_number_pos_lt_lemma: assumes \: "valid_path \" "z \ path_image \" and e: "0 < e" and ge: "\x. \0 < x; x < 1\ \ e \ Im (vector_derivative \ (at x) / (\ x - z))" shows "0 < Re(winding_number \ z)" proof - let ?vd = "\x. 1 / (\ x - z) * vector_derivative \ (at x)" let ?int = "\z. contour_integral \ (\w. 1 / (w - z))" have "e \ Im (contour_integral \ (\w. 1 / (w - z)))" proof (rule has_integral_component_le [of \ "\x. \*e" "\*e" "{0..1}", simplified]) have "((\a. 1 / (a - z)) has_contour_integral contour_integral \ (\w. 1 / (w - z))) \" thm has_integral_component_le [of \ "\x. \*e" "\*e" "{0..1}", simplified] using \ by (simp add: contour_integrable_inversediff has_contour_integral_integral) then have hi: "(?vd has_integral ?int z) (cbox 0 1)" using has_contour_integral by auto show "((\x. if 0 < x \ x < 1 then ?vd x else \ * e) has_integral ?int z) {0..1}" by (rule has_integral_spike_interior [OF hi, simplified box_real]) (use e in simp) show "\x. 0 \ x \ x \ 1 \ e \ Im (if 0 < x \ x < 1 then ?vd x else \ * e)" by (simp add: ge) qed (use has_integral_const_real [of _ 0 1] in auto) with e show ?thesis by (simp add: Re_winding_number [OF \] field_simps) qed lemma winding_number_pos_lt: assumes \: "valid_path \" "z \ path_image \" and e: "0 < e" and ge: "\x. \0 < x; x < 1\ \ e \ Im (vector_derivative \ (at x) * cnj(\ x - z))" shows "0 < Re (winding_number \ z)" proof - have bm: "bounded ((\w. w - z) ` (path_image \))" using bounded_translation [of _ "-z"] \ by (simp add: bounded_valid_path_image) then obtain B where B: "B > 0" and Bno: "\x. x \ (\w. w - z) ` (path_image \) \ norm x \ B" using bounded_pos [THEN iffD1, OF bm] by blast { fix x::real assume x: "0 < x" "x < 1" then have B2: "cmod (\ x - z)^2 \ B^2" using Bno [of "\ x - z"] by (simp add: path_image_def power2_eq_square mult_mono') with x have "\ x \ z" using \ using path_image_def by fastforce then have "e / B\<^sup>2 \ e / (cmod (\ x - z))\<^sup>2" using B B2 e by (auto simp: divide_left_mono) also have "... \ Im (vector_derivative \ (at x) * cnj (\ x - z)) / (cmod (\ x - z))\<^sup>2" using ge [OF x] by (auto simp: divide_right_mono) finally have "e / B\<^sup>2 \ Im (vector_derivative \ (at x) * cnj (\ x - z)) / (cmod (\ x - z))\<^sup>2" . then have "e / B\<^sup>2 \ Im (vector_derivative \ (at x) / (\ x - z))" by (simp add: complex_div_cnj [of _ "\ x - z" for x] del: complex_cnj_diff times_complex.sel) } note * = this show ?thesis using e B by (simp add: * winding_number_pos_lt_lemma [OF \, of "e/B^2"]) qed subsection\The winding number is an integer\ text\Proof from the book Complex Analysis by Lars V. Ahlfors, Chapter 4, section 2.1, Also on page 134 of Serge Lang's book with the name title, etc.\ lemma exp_fg: fixes z::complex assumes g: "(g has_vector_derivative g') (at x within s)" and f: "(f has_vector_derivative (g' / (g x - z))) (at x within s)" and z: "g x \ z" shows "((\x. exp(-f x) * (g x - z)) has_vector_derivative 0) (at x within s)" proof - have *: "(exp \ (\x. (- f x)) has_vector_derivative - (g' / (g x - z)) * exp (- f x)) (at x within s)" using assms unfolding has_vector_derivative_def scaleR_conv_of_real by (auto intro!: derivative_eq_intros) show ?thesis using z by (auto intro!: derivative_eq_intros * [unfolded o_def] g) qed lemma winding_number_exp_integral: fixes z::complex assumes \: "\ piecewise_C1_differentiable_on {a..b}" and ab: "a \ b" and z: "z \ \ ` {a..b}" shows "(\x. vector_derivative \ (at x) / (\ x - z)) integrable_on {a..b}" (is "?thesis1") "exp (- (integral {a..b} (\x. vector_derivative \ (at x) / (\ x - z)))) * (\ b - z) = \ a - z" (is "?thesis2") proof - let ?D\ = "\x. vector_derivative \ (at x)" have [simp]: "\x. a \ x \ x \ b \ \ x \ z" using z by force have con_g: "continuous_on {a..b} \" using \ by (simp add: piecewise_C1_differentiable_on_def) obtain k where fink: "finite k" and g_C1_diff: "\ C1_differentiable_on ({a..b} - k)" using \ by (force simp: piecewise_C1_differentiable_on_def) have \: "open ({a<..finite k\ by (simp add: finite_imp_closed open_Diff) moreover have "{a<.. {a..b} - k" by force ultimately have g_diff_at: "\x. \x \ k; x \ {a<.. \ \ differentiable at x" by (metis Diff_iff differentiable_on_subset C1_diff_imp_diff [OF g_C1_diff] differentiable_on_def at_within_open) { fix w assume "w \ z" have "continuous_on (ball w (cmod (w - z))) (\w. 1 / (w - z))" by (auto simp: dist_norm intro!: continuous_intros) moreover have "\x. cmod (w - x) < cmod (w - z) \ \f'. ((\w. 1 / (w - z)) has_field_derivative f') (at x)" by (auto simp: intro!: derivative_eq_intros) ultimately have "\h. \y. norm(y - w) < norm(w - z) \ (h has_field_derivative 1/(y - z)) (at y)" using holomorphic_convex_primitive [of "ball w (norm(w - z))" "{}" "\w. 1/(w - z)"] by (force simp: field_differentiable_def Ball_def dist_norm at_within_open_NO_MATCH norm_minus_commute) } then obtain h where h: "\w y. w \ z \ norm(y - w) < norm(w - z) \ (h w has_field_derivative 1/(y - z)) (at y)" by meson have exy: "\y. ((\x. inverse (\ x - z) * ?D\ x) has_integral y) {a..b}" unfolding integrable_on_def [symmetric] proof (rule contour_integral_local_primitive_any [OF piecewise_C1_imp_differentiable [OF \]]) show "\d h. 0 < d \ (\y. cmod (y - w) < d \ (h has_field_derivative inverse (y - z))(at y within - {z}))" if "w \ - {z}" for w using that inverse_eq_divide has_field_derivative_at_within h by (metis Compl_insert DiffD2 insertCI right_minus_eq zero_less_norm_iff) qed simp have vg_int: "(\x. ?D\ x / (\ x - z)) integrable_on {a..b}" unfolding box_real [symmetric] divide_inverse_commute by (auto intro!: exy integrable_subinterval simp add: integrable_on_def ab) with ab show ?thesis1 by (simp add: divide_inverse_commute integral_def integrable_on_def) { fix t assume t: "t \ {a..b}" have cball: "continuous_on (ball (\ t) (dist (\ t) z)) (\x. inverse (x - z))" using z by (auto intro!: continuous_intros simp: dist_norm) have icd: "\x. cmod (\ t - x) < cmod (\ t - z) \ (\w. inverse (w - z)) field_differentiable at x" unfolding field_differentiable_def by (force simp: intro!: derivative_eq_intros) obtain h where h: "\x. cmod (\ t - x) < cmod (\ t - z) \ (h has_field_derivative inverse (x - z)) (at x within {y. cmod (\ t - y) < cmod (\ t - z)})" using holomorphic_convex_primitive [where f = "\w. inverse(w - z)", OF convex_ball finite.emptyI cball icd] by simp (auto simp: ball_def dist_norm that) have "exp (- (integral {a..t} (\x. ?D\ x / (\ x - z)))) * (\ t - z) =\ a - z" proof (rule has_derivative_zero_unique_strong_interval [of "{a,b} \ k" a b]) show "continuous_on {a..b} (\b. exp (- integral {a..b} (\x. ?D\ x / (\ x - z))) * (\ b - z))" by (auto intro!: continuous_intros con_g indefinite_integral_continuous_1 [OF vg_int]) show "((\b. exp (- integral {a..b} (\x. ?D\ x / (\ x - z))) * (\ b - z)) has_derivative (\h. 0)) (at x within {a..b})" if "x \ {a..b} - ({a, b} \ k)" for x proof - have x: "x \ k" "a < x" "x < b" using that by auto then have "x \ interior ({a..b} - k)" using open_subset_interior [OF \] by fastforce then have con: "isCont ?D\ x" using g_C1_diff x by (auto simp: C1_differentiable_on_eq intro: continuous_on_interior) then have con_vd: "continuous (at x within {a..b}) (\x. ?D\ x)" by (rule continuous_at_imp_continuous_within) have gdx: "\ differentiable at x" using x by (simp add: g_diff_at) then obtain d where d: "(\ has_derivative (\x. x *\<^sub>R d)) (at x)" by (auto simp add: differentiable_iff_scaleR) show "((\c. exp (- integral {a..c} (\x. ?D\ x / (\ x - z))) * (\ c - z)) has_derivative (\h. 0)) (at x within {a..b})" proof (rule exp_fg [unfolded has_vector_derivative_def, simplified]) show "(\ has_derivative (\c. c *\<^sub>R d)) (at x within {a..b})" using d by (blast intro: has_derivative_at_withinI) have "((\x. integral {a..x} (\x. ?D\ x / (\ x - z))) has_vector_derivative d / (\ x - z)) (at x within {a..b})" proof (rule has_vector_derivative_eq_rhs [OF integral_has_vector_derivative_continuous_at [where S = "{}", simplified]]) show "continuous (at x within {a..b}) (\x. vector_derivative \ (at x) / (\ x - z))" using continuous_at_imp_continuous_at_within differentiable_imp_continuous_within gdx x by (intro con_vd continuous_intros) (force+) show "vector_derivative \ (at x) / (\ x - z) = d / (\ x - z)" using d vector_derivative_at by (simp add: vector_derivative_at has_vector_derivative_def) qed (use x vg_int in auto) then show "((\x. integral {a..x} (\x. ?D\ x / (\ x - z))) has_derivative (\c. c *\<^sub>R (d / (\ x - z)))) (at x within {a..b})" by (auto simp: has_vector_derivative_def) qed (use x in auto) qed qed (use fink t in auto) } with ab show ?thesis2 by (simp add: divide_inverse_commute integral_def) qed lemma winding_number_exp_2pi: "\path p; z \ path_image p\ \ pathfinish p - z = exp (2 * pi * \ * winding_number p z) * (pathstart p - z)" using winding_number [of p z 1] unfolding valid_path_def path_image_def pathstart_def pathfinish_def winding_number_prop_def by (force dest: winding_number_exp_integral(2) [of _ 0 1 z] simp: field_simps contour_integral_integral exp_minus) lemma integer_winding_number_eq: assumes \: "path \" and z: "z \ path_image \" shows "winding_number \ z \ \ \ pathfinish \ = pathstart \" proof - obtain p where p: "valid_path p" "z \ path_image p" "pathstart p = pathstart \" "pathfinish p = pathfinish \" and eq: "contour_integral p (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number \ z" using winding_number [OF assms, of 1] unfolding winding_number_prop_def by auto then have wneq: "winding_number \ z = winding_number p z" using eq winding_number_valid_path by force have iff: "(winding_number \ z \ \) \ (exp (contour_integral p (\w. 1 / (w - z))) = 1)" using eq by (simp add: exp_eq_1 complex_is_Int_iff) have "\ 0 \ z" by (metis pathstart_def pathstart_in_path_image z) then have "exp (contour_integral p (\w. 1 / (w - z))) = (\ 1 - z) / (\ 0 - z)" using p winding_number_exp_integral(2) [of p 0 1 z] by (simp add: valid_path_def path_defs contour_integral_integral exp_minus field_split_simps) then have "winding_number p z \ \ \ pathfinish p = pathstart p" using p wneq iff by (auto simp: path_defs) then show ?thesis using p eq by (auto simp: winding_number_valid_path) qed theorem integer_winding_number: "\path \; pathfinish \ = pathstart \; z \ path_image \\ \ winding_number \ z \ \" by (metis integer_winding_number_eq) text\If the winding number's magnitude is at least one, then the path must contain points in every direction.*) We can thus bound the winding number of a path that doesn't intersect a given ray. \ lemma winding_number_pos_meets: fixes z::complex assumes \: "valid_path \" and z: "z \ path_image \" and 1: "Re (winding_number \ z) \ 1" and w: "w \ z" shows "\a::real. 0 < a \ z + of_real a * (w - z) \ path_image \" proof - have [simp]: "\x. 0 \ x \ x \ 1 \ \ x \ z" using z by (auto simp: path_image_def) have [simp]: "z \ \ ` {0..1}" using path_image_def z by auto have gpd: "\ piecewise_C1_differentiable_on {0..1}" using \ valid_path_def by blast define r where "r = (w - z) / (\ 0 - z)" have [simp]: "r \ 0" using w z by (auto simp: r_def) have cont: "continuous_on {0..1} (\x. Im (integral {0..x} (\x. vector_derivative \ (at x) / (\ x - z))))" by (intro continuous_intros indefinite_integral_continuous_1 winding_number_exp_integral [OF gpd]; simp) have "Arg2pi r \ 2*pi" by (simp add: Arg2pi less_eq_real_def) also have "\ \ Im (integral {0..1} (\x. vector_derivative \ (at x) / (\ x - z)))" using 1 by (simp add: winding_number_valid_path [OF \ z] contour_integral_integral Complex.Re_divide field_simps power2_eq_square) finally have "Arg2pi r \ Im (integral {0..1} (\x. vector_derivative \ (at x) / (\ x - z)))" . then have "\t. t \ {0..1} \ Im(integral {0..t} (\x. vector_derivative \ (at x)/(\ x - z))) = Arg2pi r" by (simp add: Arg2pi_ge_0 cont IVT') then obtain t where t: "t \ {0..1}" and eqArg: "Im (integral {0..t} (\x. vector_derivative \ (at x)/(\ x - z))) = Arg2pi r" by blast define i where "i = integral {0..t} (\x. vector_derivative \ (at x) / (\ x - z))" have gpdt: "\ piecewise_C1_differentiable_on {0..t}" by (metis atLeastAtMost_iff atLeastatMost_subset_iff order_refl piecewise_C1_differentiable_on_subset gpd t) have "exp (- i) * (\ t - z) = \ 0 - z" unfolding i_def proof (rule winding_number_exp_integral [OF gpdt]) show "z \ \ ` {0..t}" using t z unfolding path_image_def by force qed (use t in auto) then have *: "\ t - z = exp i * (\ 0 - z)" by (simp add: exp_minus field_simps) then have "(w - z) = r * (\ 0 - z)" by (simp add: r_def) moreover have "z + exp (Re i) * (exp (\ * Im i) * (\ 0 - z)) = \ t" using * by (simp add: exp_eq_polar field_simps) moreover have "Arg2pi r = Im i" using eqArg by (simp add: i_def) ultimately have "z + complex_of_real (exp (Re i)) * (w - z) / complex_of_real (cmod r) = \ t" using Complex_Transcendental.Arg2pi_eq [of r] \r \ 0\ by (metis mult.left_commute nonzero_mult_div_cancel_left norm_eq_zero of_real_0 of_real_eq_iff times_divide_eq_left) with t show ?thesis by (rule_tac x="exp(Re i) / norm r" in exI) (auto simp: path_image_def) qed lemma winding_number_big_meets: fixes z::complex assumes \: "valid_path \" and z: "z \ path_image \" and "\Re (winding_number \ z)\ \ 1" and w: "w \ z" shows "\a::real. 0 < a \ z + of_real a * (w - z) \ path_image \" proof - { assume "Re (winding_number \ z) \ - 1" then have "Re (winding_number (reversepath \) z) \ 1" by (simp add: \ valid_path_imp_path winding_number_reversepath z) moreover have "valid_path (reversepath \)" using \ valid_path_imp_reverse by auto moreover have "z \ path_image (reversepath \)" by (simp add: z) ultimately have "\a::real. 0 < a \ z + of_real a * (w - z) \ path_image (reversepath \)" using winding_number_pos_meets w by blast then have ?thesis by simp } then show ?thesis using assms by (simp add: abs_if winding_number_pos_meets split: if_split_asm) qed lemma winding_number_less_1: fixes z::complex shows "\valid_path \; z \ path_image \; w \ z; \a::real. 0 < a \ z + of_real a * (w - z) \ path_image \\ \ Re(winding_number \ z) < 1" by (auto simp: not_less dest: winding_number_big_meets) text\One way of proving that WN=1 for a loop.\ lemma winding_number_eq_1: fixes z::complex assumes \: "valid_path \" and z: "z \ path_image \" and loop: "pathfinish \ = pathstart \" and 0: "0 < Re(winding_number \ z)" and 2: "Re(winding_number \ z) < 2" shows "winding_number \ z = 1" proof - have "winding_number \ z \ Ints" by (simp add: \ integer_winding_number loop valid_path_imp_path z) then show ?thesis using 0 2 by (auto simp: Ints_def) qed subsection\Continuity of winding number and invariance on connected sets\ theorem continuous_at_winding_number: fixes z::complex assumes \: "path \" and z: "z \ path_image \" shows "continuous (at z) (winding_number \)" proof - obtain e where "e>0" and cbg: "cball z e \ - path_image \" using open_contains_cball [of "- path_image \"] z by (force simp: closed_def [symmetric] closed_path_image [OF \]) then have ppag: "path_image \ \ - cball z (e/2)" by (force simp: cball_def dist_norm) have oc: "open (- cball z (e/2))" by (simp add: closed_def [symmetric]) obtain d where "d>0" and pi_eq: "\h1 h2. \valid_path h1; valid_path h2; (\t\{0..1}. cmod (h1 t - \ t) < d \ cmod (h2 t - \ t) < d); pathstart h2 = pathstart h1; pathfinish h2 = pathfinish h1\ \ path_image h1 \ - cball z (e/2) \ path_image h2 \ - cball z (e/2) \ (\f. f holomorphic_on - cball z (e/2) \ contour_integral h2 f = contour_integral h1 f)" using contour_integral_nearby_ends [OF oc \ ppag] by metis obtain p where "valid_path p" "z \ path_image p" and p: "pathstart p = pathstart \" "pathfinish p = pathfinish \" and pg: "\t. t\{0..1} \ cmod (\ t - p t) < min d e/2" and pi: "contour_integral p (\x. 1 / (x - z)) = complex_of_real (2 * pi) * \ * winding_number \ z" using winding_number [OF \ z, of "min d e/2"] \d>0\ \e>0\ by (auto simp: winding_number_prop_def) { fix w assume d2: "cmod (w - z) < d/2" and e2: "cmod (w - z) < e/2" have wnotp: "w \ path_image p" proof (clarsimp simp add: path_image_def) show False if w: "w = p x" and "0 \ x" "x \ 1" for x proof - have "cmod (\ x - p x) < min d e/2" using pg that by auto then have "cmod (z - \ x) < e" by (metis e2 less_divide_eq_numeral1(1) min_less_iff_conj norm_minus_commute norm_triangle_half_l w) then show ?thesis using cbg that by (auto simp add: path_image_def cball_def dist_norm less_eq_real_def) qed qed have wnotg: "w \ path_image \" using cbg e2 \e>0\ by (force simp: dist_norm norm_minus_commute) { fix k::real assume k: "k>0" then obtain q where q: "valid_path q" "w \ path_image q" "pathstart q = pathstart \ \ pathfinish q = pathfinish \" and qg: "\t. t \ {0..1} \ cmod (\ t - q t) < min k (min d e) / 2" and qi: "contour_integral q (\u. 1 / (u - w)) = complex_of_real (2 * pi) * \ * winding_number \ w" using winding_number [OF \ wnotg, of "min k (min d e) / 2"] \d>0\ \e>0\ k by (force simp: min_divide_distrib_right winding_number_prop_def) moreover have "\t. t \ {0..1} \ cmod (q t - \ t) < d \ cmod (p t - \ t) < d" using pg qg \0 < d\ by (fastforce simp add: norm_minus_commute) moreover have "(\u. 1 / (u-w)) holomorphic_on - cball z (e/2)" using e2 by (auto simp: dist_norm norm_minus_commute intro!: holomorphic_intros) ultimately have "contour_integral p (\u. 1 / (u - w)) = contour_integral q (\u. 1 / (u - w))" by (metis p \valid_path p\ pi_eq) then have "contour_integral p (\x. 1 / (x - w)) = complex_of_real (2 * pi) * \ * winding_number \ w" by (simp add: pi qi) } note pip = this have "path p" by (simp add: \valid_path p\ valid_path_imp_path) moreover have "\e. e > 0 \ winding_number_prop p w e p (winding_number \ w)" by (simp add: \valid_path p\ pip winding_number_prop_def wnotp) ultimately have "winding_number p w = winding_number \ w" using winding_number_unique wnotp by blast } note wnwn = this obtain pe where "pe>0" and cbp: "cball z (3 / 4 * pe) \ - path_image p" using \valid_path p\ \z \ path_image p\ open_contains_cball [of "- path_image p"] by (force simp: closed_def [symmetric] closed_path_image [OF valid_path_imp_path]) obtain L where "L>0" and L: "\f B. \f holomorphic_on - cball z (3 / 4 * pe); \z \ - cball z (3 / 4 * pe). cmod (f z) \ B\ \ cmod (contour_integral p f) \ L * B" using contour_integral_bound_exists [of "- cball z (3/4*pe)" p] cbp \valid_path p\ by blast { fix e::real and w::complex assume e: "0 < e" and w: "cmod (w - z) < pe/4" "cmod (w - z) < e * pe\<^sup>2 / (8 * L)" then have [simp]: "w \ path_image p" using cbp p(2) \0 < pe\ by (force simp: dist_norm norm_minus_commute path_image_def cball_def) have [simp]: "contour_integral p (\x. 1/(x - w)) - contour_integral p (\x. 1/(x - z)) = contour_integral p (\x. 1/(x - w) - 1/(x - z))" by (simp add: \valid_path p\ \z \ path_image p\ contour_integrable_inversediff contour_integral_diff) { fix x assume pe: "3/4 * pe < cmod (z - x)" have "cmod (w - x) < pe/4 + cmod (z - x)" by (meson add_less_cancel_right norm_diff_triangle_le order_refl order_trans_rules(21) w(1)) then have wx: "cmod (w - x) < 4/3 * cmod (z - x)" using pe by simp have "cmod (z - x) \ cmod (z - w) + cmod (w - x)" using norm_diff_triangle_le by blast also have "\ < pe/4 + cmod (w - x)" using w by (simp add: norm_minus_commute) finally have "pe/2 < cmod (w - x)" using pe by auto then have pe_less: "(pe/2)^2 < cmod (w - x) ^ 2" by (simp add: \0 < pe\ less_eq_real_def power_strict_mono) then have pe2: "pe^2 < 4 * cmod (w - x) ^ 2" by (simp add: power_divide) have "8 * L * cmod (w - z) < e * pe\<^sup>2" using w \L>0\ by (simp add: field_simps) also have "\ < e * 4 * cmod (w - x) * cmod (w - x)" using pe2 \e>0\ by (simp add: power2_eq_square) also have "\ < e * 4 * cmod (w - x) * (4/3 * cmod (z - x))" using \0 < pe\ pe_less e less_eq_real_def wx by fastforce finally have "L * cmod (w - z) < 2/3 * e * cmod (w - x) * cmod (z - x)" by simp also have "\ \ e * cmod (w - x) * cmod (z - x)" using e by simp finally have Lwz: "L * cmod (w - z) < e * cmod (w - x) * cmod (z - x)" . have "L * cmod (1 / (x - w) - 1 / (x - z)) \ e" proof (cases "x=z \ x=w") case True with pe \pe>0\ w \L>0\ show ?thesis by (force simp: norm_minus_commute) next case False with wx w(2) \L>0\ pe pe2 Lwz show ?thesis by (auto simp: divide_simps mult_less_0_iff norm_minus_commute norm_divide norm_mult power2_eq_square) qed } note L_cmod_le = this let ?f = "(\x. 1 / (x - w) - 1 / (x - z))" have "cmod (contour_integral p ?f) \ L * (e * pe\<^sup>2 / L / 4 * (inverse (pe / 2))\<^sup>2)" proof (rule L) show "?f holomorphic_on - cball z (3 / 4 * pe)" using \pe>0\ w by (force simp: dist_norm norm_minus_commute intro!: holomorphic_intros) show " \u \- cball z (3 / 4 * pe). cmod (?f u) \ e * pe\<^sup>2 / L / 4 * (inverse (pe / 2))\<^sup>2" using \pe>0\ w \L>0\ by (auto simp: cball_def dist_norm field_simps L_cmod_le simp del: less_divide_eq_numeral1 le_divide_eq_numeral1) qed also have "\ < 2*e" using \L>0\ e by (force simp: field_simps) finally have "cmod (winding_number p w - winding_number p z) < e" using pi_ge_two e by (force simp: winding_number_valid_path \valid_path p\ \z \ path_image p\ field_simps norm_divide norm_mult intro: less_le_trans) } note cmod_wn_diff = this have "isCont (winding_number p) z" proof (clarsimp simp add: continuous_at_eps_delta) fix e::real assume "e>0" show "\d>0. \x'. dist x' z < d \ dist (winding_number p x') (winding_number p z) < e" using \pe>0\ \L>0\ \e>0\ by (rule_tac x="min (pe/4) (e/2*pe^2/L/4)" in exI) (simp add: dist_norm cmod_wn_diff) qed then show ?thesis apply (rule continuous_transform_within [where d = "min d e/2"]) apply (auto simp: \d>0\ \e>0\ dist_norm wnwn) done qed corollary continuous_on_winding_number: "path \ \ continuous_on (- path_image \) (\w. winding_number \ w)" by (simp add: continuous_at_imp_continuous_on continuous_at_winding_number) subsection\<^marker>\tag unimportant\ \The winding number is constant on a connected region\ lemma winding_number_constant: assumes \: "path \" and loop: "pathfinish \ = pathstart \" and cs: "connected S" and sg: "S \ path_image \ = {}" shows "winding_number \ constant_on S" proof - have *: "1 \ cmod (winding_number \ y - winding_number \ z)" if ne: "winding_number \ y \ winding_number \ z" and "y \ S" "z \ S" for y z proof - have "winding_number \ y \ \" "winding_number \ z \ \" using that integer_winding_number [OF \ loop] sg \y \ S\ by auto with ne show ?thesis by (auto simp: Ints_def simp flip: of_int_diff) qed have cont: "continuous_on S (\w. winding_number \ w)" using continuous_on_winding_number [OF \] sg by (meson continuous_on_subset disjoint_eq_subset_Compl) show ?thesis using "*" zero_less_one by (blast intro: continuous_discrete_range_constant [OF cs cont]) qed lemma winding_number_eq: "\path \; pathfinish \ = pathstart \; w \ S; z \ S; connected S; S \ path_image \ = {}\ \ winding_number \ w = winding_number \ z" using winding_number_constant by (metis constant_on_def) lemma open_winding_number_levelsets: assumes \: "path \" and loop: "pathfinish \ = pathstart \" shows "open {z. z \ path_image \ \ winding_number \ z = k}" proof (clarsimp simp: open_dist) fix z assume z: "z \ path_image \" and k: "k = winding_number \ z" have "open (- path_image \)" by (simp add: closed_path_image \ open_Compl) then obtain e where "e>0" "ball z e \ - path_image \" using open_contains_ball [of "- path_image \"] z by blast then show "\e>0. \y. dist y z < e \ y \ path_image \ \ winding_number \ y = winding_number \ z" using \e>0\ by (force simp: norm_minus_commute dist_norm intro: winding_number_eq [OF assms, where S = "ball z e"]) qed subsection\Winding number is zero "outside" a curve\ proposition winding_number_zero_in_outside: assumes \: "path \" and loop: "pathfinish \ = pathstart \" and z: "z \ outside (path_image \)" shows "winding_number \ z = 0" proof - obtain B::real where "0 < B" and B: "path_image \ \ ball 0 B" using bounded_subset_ballD [OF bounded_path_image [OF \]] by auto obtain w::complex where w: "w \ ball 0 (B + 1)" by (metis abs_of_nonneg le_less less_irrefl mem_ball_0 norm_of_real) have "- ball 0 (B + 1) \ outside (path_image \)" using B subset_ball by (intro outside_subset_convex) auto then have wout: "w \ outside (path_image \)" using w by blast moreover have "winding_number \ constant_on outside (path_image \)" using winding_number_constant [OF \ loop, of "outside(path_image \)"] connected_outside by (metis DIM_complex bounded_path_image dual_order.refl \ outside_no_overlap) ultimately have "winding_number \ z = winding_number \ w" by (metis (no_types, opaque_lifting) constant_on_def z) also have "\ = 0" proof - have wnot: "w \ path_image \" using wout by (simp add: outside_def) { fix e::real assume "0" "pathfinish p = pathfinish \" and pg1: "(\t. \0 \ t; t \ 1\ \ cmod (p t - \ t) < 1)" and pge: "(\t. \0 \ t; t \ 1\ \ cmod (p t - \ t) < e)" using path_approx_polynomial_function [OF \, of "min 1 e"] \e>0\ by (metis atLeastAtMost_iff min_less_iff_conj zero_less_one) have "\p. valid_path p \ w \ path_image p \ pathstart p = pathstart \ \ pathfinish p = pathfinish \ \ (\t\{0..1}. cmod (\ t - p t) < e) \ contour_integral p (\wa. 1 / (wa - w)) = 0" proof (intro exI conjI) have "\x. \0 \ x; x \ 1\ \ cmod (p x) < B + 1" using B unfolding image_subset_iff path_image_def by (meson add_strict_mono atLeastAtMost_iff le_less_trans mem_ball_0 norm_triangle_sub pg1) then have pip: "path_image p \ ball 0 (B + 1)" by (auto simp add: path_image_def dist_norm ball_def) then show "w \ path_image p" using w by blast show vap: "valid_path p" by (simp add: p(1) valid_path_polynomial_function) show "\t\{0..1}. cmod (\ t - p t) < e" by (metis atLeastAtMost_iff norm_minus_commute pge) show "contour_integral p (\wa. 1 / (wa - w)) = 0" proof (rule contour_integral_unique [OF Cauchy_theorem_convex_simple [OF _ convex_ball [of 0 "B+1"]]]) have "\z. cmod z < B + 1 \ z \ w" using mem_ball_0 w by blast then show "(\z. 1 / (z - w)) holomorphic_on ball 0 (B + 1)" by (intro holomorphic_intros; simp add: dist_norm) qed (use p vap pip loop in auto) qed (use p in auto) } then show ?thesis by (auto intro: winding_number_unique [OF \] simp add: winding_number_prop_def wnot) qed finally show ?thesis . qed corollary\<^marker>\tag unimportant\ winding_number_zero_const: "a \ z \ winding_number (\t. a) z = 0" by (rule winding_number_zero_in_outside) (auto simp: pathfinish_def pathstart_def path_polynomial_function) corollary\<^marker>\tag unimportant\ winding_number_zero_outside: "\path \; convex s; pathfinish \ = pathstart \; z \ s; path_image \ \ s\ \ winding_number \ z = 0" by (meson convex_in_outside outside_mono subsetCE winding_number_zero_in_outside) lemma winding_number_zero_at_infinity: assumes \: "path \" and loop: "pathfinish \ = pathstart \" shows "\B. \z. B \ norm z \ winding_number \ z = 0" proof - obtain B::real where "0 < B" and B: "path_image \ \ ball 0 B" using bounded_subset_ballD [OF bounded_path_image [OF \]] by auto have "winding_number \ z = 0" if "B + 1 \ cmod z" for z proof (rule winding_number_zero_outside [OF \ convex_cball loop]) show "z \ cball 0 B" using that by auto show "path_image \ \ cball 0 B" using B order.trans by blast qed then show ?thesis by metis qed lemma winding_number_zero_point: "\path \; convex S; pathfinish \ = pathstart \; open S; path_image \ \ S\ \ \z. z \ S \ winding_number \ z = 0" using outside_compact_in_open [of "path_image \" S] path_image_nonempty winding_number_zero_in_outside by (fastforce simp add: compact_path_image) text\If a path winds round a set, it winds rounds its inside.\ lemma winding_number_around_inside: assumes \: "path \" and loop: "pathfinish \ = pathstart \" and cls: "closed S" and cos: "connected S" and S_disj: "S \ path_image \ = {}" and z: "z \ S" and wn_nz: "winding_number \ z \ 0" and w: "w \ S \ inside S" shows "winding_number \ w = winding_number \ z" proof - have ssb: "S \ inside(path_image \)" proof fix x :: complex assume "x \ S" hence "x \ path_image \" by (meson disjoint_iff_not_equal S_disj) thus "x \ inside (path_image \)" by (metis Compl_iff S_disj UnE \ \x \ S\ cos inside_outside loop winding_number_eq winding_number_zero_in_outside wn_nz z) qed show ?thesis proof (rule winding_number_eq [OF \ loop w]) show "z \ S \ inside S" using z by blast show "connected (S \ inside S)" by (simp add: cls connected_with_inside cos) show "(S \ inside S) \ path_image \ = {}" unfolding disjoint_iff Un_iff by (metis ComplD UnI1 \ cls compact_path_image connected_path_image inside_Un_outside inside_inside_compact_connected ssb subsetD) qed qed text\Bounding a WN by 1/2 for a path and point in opposite halfspaces.\ lemma winding_number_subpath_continuous: assumes \: "valid_path \" and z: "z \ path_image \" shows "continuous_on {0..1} (\x. winding_number(subpath 0 x \) z)" proof (rule continuous_on_eq) let ?f = "\x. integral {0..x} (\t. vector_derivative \ (at t) / (\ t - z))" show "continuous_on {0..1} (\x. 1 / (2 * pi * \) * ?f x)" proof (intro indefinite_integral_continuous_1 winding_number_exp_integral continuous_intros) show "\ piecewise_C1_differentiable_on {0..1}" using \ valid_path_def by blast qed (use path_image_def z in auto) show "1 / (2 * pi * \) * ?f x = winding_number (subpath 0 x \) z" if x: "x \ {0..1}" for x proof - have "1 / (2*pi*\) * ?f x = 1 / (2*pi*\) * contour_integral (subpath 0 x \) (\w. 1/(w - z))" using assms x by (simp add: contour_integral_subcontour_integral [OF contour_integrable_inversediff]) also have "\ = winding_number (subpath 0 x \) z" proof (subst winding_number_valid_path) show "z \ path_image (subpath 0 x \)" using assms x atLeastAtMost_iff path_image_subpath_subset by force qed (use assms x valid_path_subpath in \force+\) finally show ?thesis . qed qed lemma winding_number_ivt_pos: assumes \: "valid_path \" and z: "z \ path_image \" and "0 \ w" "w \ Re(winding_number \ z)" shows "\t \ {0..1}. Re(winding_number(subpath 0 t \) z) = w" proof - have "continuous_on {0..1} (\x. winding_number (subpath 0 x \) z)" using \ winding_number_subpath_continuous z by blast moreover have "Re (winding_number (subpath 0 0 \) z) \ w" "w \ Re (winding_number (subpath 0 1 \) z)" using assms by (auto simp: path_image_def image_def) ultimately show ?thesis using ivt_increasing_component_on_1[of 0 1, where ?k = "1"] by force qed lemma winding_number_ivt_neg: assumes \: "valid_path \" and z: "z \ path_image \" and "Re(winding_number \ z) \ w" "w \ 0" shows "\t \ {0..1}. Re(winding_number(subpath 0 t \) z) = w" proof - have "continuous_on {0..1} (\x. winding_number (subpath 0 x \) z)" using \ winding_number_subpath_continuous z by blast moreover have "Re (winding_number (subpath 0 0 \) z) \ w" "w \ Re (winding_number (subpath 0 1 \) z)" using assms by (auto simp: path_image_def image_def) ultimately show ?thesis using ivt_decreasing_component_on_1[of 0 1, where ?k = "1"] by force qed lemma winding_number_ivt_abs: assumes \: "valid_path \" and z: "z \ path_image \" and "0 \ w" "w \ \Re(winding_number \ z)\" shows "\t \ {0..1}. \Re (winding_number (subpath 0 t \) z)\ = w" using assms winding_number_ivt_pos [of \ z w] winding_number_ivt_neg [of \ z "-w"] by force lemma winding_number_lt_half_lemma: assumes \: "valid_path \" and z: "z \ path_image \" and az: "a \ z \ b" and pag: "path_image \ \ {w. a \ w > b}" shows "Re(winding_number \ z) < 1/2" proof - { assume "Re(winding_number \ z) \ 1/2" then obtain t::real where t: "0 \ t" "t \ 1" and sub12: "Re (winding_number (subpath 0 t \) z) = 1/2" using winding_number_ivt_pos [OF \ z, of "1/2"] by auto have gt: "\ t - z = - (of_real (exp (- (2 * pi * Im (winding_number (subpath 0 t \) z)))) * (\ 0 - z))" using winding_number_exp_2pi [of "subpath 0 t \" z] apply (simp add: t \ valid_path_imp_path) using closed_segment_eq_real_ivl path_image_def t z by (fastforce simp: path_image_subpath Euler sub12) have "b < a \ \ 0" proof - have "\ 0 \ {c. b < a \ c}" by (metis (no_types) pag atLeastAtMost_iff image_subset_iff order_refl path_image_def zero_le_one) thus ?thesis by blast qed moreover have "b < a \ \ t" by (metis atLeastAtMost_iff image_eqI mem_Collect_eq pag path_image_def subset_iff t) ultimately have "0 < a \ (\ 0 - z)" "0 < a \ (\ t - z)" using az by (simp add: inner_diff_right)+ then have False by (simp add: gt inner_mult_right mult_less_0_iff) } then show ?thesis by force qed lemma winding_number_lt_half: assumes "valid_path \" "a \ z \ b" "path_image \ \ {w. a \ w > b}" shows "\Re (winding_number \ z)\ < 1/2" proof - have "z \ path_image \" using assms by auto with assms have "Re (winding_number \ z) < 0 \ - Re (winding_number \ z) < 1/2" by (metis complex_inner_1_right winding_number_lt_half_lemma [OF valid_path_imp_reverse, of \ z a b] winding_number_reversepath valid_path_imp_path inner_minus_left path_image_reversepath) with assms show ?thesis using \z \ path_image \\ winding_number_lt_half_lemma by fastforce qed lemma winding_number_le_half: assumes \: "valid_path \" and z: "z \ path_image \" and anz: "a \ 0" and azb: "a \ z \ b" and pag: "path_image \ \ {w. a \ w \ b}" shows "\Re (winding_number \ z)\ \ 1/2" proof - { assume wnz_12: "\Re (winding_number \ z)\ > 1/2" have "isCont (winding_number \) z" by (metis continuous_at_winding_number valid_path_imp_path \ z) then obtain d where "d>0" and d: "\x'. dist x' z < d \ dist (winding_number \ x') (winding_number \ z) < \Re(winding_number \ z)\ - 1/2" using continuous_at_eps_delta wnz_12 diff_gt_0_iff_gt by blast define z' where "z' = z - (d / (2 * cmod a)) *\<^sub>R a" have "a \ z * 6 \ d * cmod a + b * 6" by (metis \0 < d\ add_increasing azb less_eq_real_def mult_nonneg_nonneg mult_right_mono norm_ge_zero norm_numeral) with anz have *: "a \ z' \ b - d / 3 * cmod a" unfolding z'_def inner_mult_right' divide_inverse by (simp add: field_split_simps algebra_simps dot_square_norm power2_eq_square) have "cmod (winding_number \ z' - winding_number \ z) < \Re (winding_number \ z)\ - 1/2" using d [of z'] anz \d>0\ by (simp add: dist_norm z'_def) then have "1/2 < \Re (winding_number \ z)\ - cmod (winding_number \ z' - winding_number \ z)" by simp then have "1/2 < \Re (winding_number \ z)\ - \Re (winding_number \ z') - Re (winding_number \ z)\" using abs_Re_le_cmod [of "winding_number \ z' - winding_number \ z"] by simp then have wnz_12': "\Re (winding_number \ z')\ > 1/2" by linarith moreover have "\Re (winding_number \ z')\ < 1/2" proof (rule winding_number_lt_half [OF \ *]) show "path_image \ \ {w. b - d / 3 * cmod a < a \ w}" using azb \d>0\ pag by (auto simp: add_strict_increasing anz field_split_simps dest!: subsetD) qed ultimately have False by simp } then show ?thesis by force qed lemma winding_number_lt_half_linepath: assumes "z \ closed_segment a b" shows "\Re (winding_number (linepath a b) z)\ < 1/2" proof - obtain u v where "u \ z \ v" and uv: "\x. x \ closed_segment a b \ inner u x > v" using separating_hyperplane_closed_point assms closed_segment convex_closed_segment less_eq_real_def by metis moreover have "path_image (linepath a b) \ {w. v < u \ w}" using in_segment(1) uv by auto ultimately show ?thesis using winding_number_lt_half by auto qed text\ Positivity of WN for a linepath.\ lemma winding_number_linepath_pos_lt: assumes "0 < Im ((b - a) * cnj (b - z))" shows "0 < Re(winding_number(linepath a b) z)" proof - have z: "z \ path_image (linepath a b)" using assms by (simp add: closed_segment_def) (force simp: algebra_simps) show ?thesis by (intro winding_number_pos_lt [OF valid_path_linepath z assms]) (simp add: linepath_def algebra_simps) qed subsection\<^marker>\tag unimportant\ \More winding number properties\ text\including the fact that it's +-1 inside a simple closed curve.\ lemma winding_number_homotopic_paths: assumes "homotopic_paths (-{z}) g h" shows "winding_number g z = winding_number h z" proof - have "path g" "path h" using homotopic_paths_imp_path [OF assms] by auto moreover have pag: "z \ path_image g" and pah: "z \ path_image h" using homotopic_paths_imp_subset [OF assms] by auto ultimately obtain d e where "d > 0" "e > 0" and d: "\p. \path p; pathstart p = pathstart g; pathfinish p = pathfinish g; \t\{0..1}. norm (p t - g t) < d\ \ homotopic_paths (-{z}) g p" and e: "\q. \path q; pathstart q = pathstart h; pathfinish q = pathfinish h; \t\{0..1}. norm (q t - h t) < e\ \ homotopic_paths (-{z}) h q" using homotopic_nearby_paths [of g "-{z}"] homotopic_nearby_paths [of h "-{z}"] by force obtain p where p: "valid_path p" "z \ path_image p" "pathstart p = pathstart g" "pathfinish p = pathfinish g" and gp_less:"\t\{0..1}. cmod (g t - p t) < d" and pap: "contour_integral p (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number g z" using winding_number [OF \path g\ pag \0 < d\] unfolding winding_number_prop_def by blast obtain q where q: "valid_path q" "z \ path_image q" "pathstart q = pathstart h" "pathfinish q = pathfinish h" and hq_less: "\t\{0..1}. cmod (h t - q t) < e" and paq: "contour_integral q (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number h z" using winding_number [OF \path h\ pah \0 < e\] unfolding winding_number_prop_def by blast have "homotopic_paths (- {z}) g p" by (simp add: d p valid_path_imp_path norm_minus_commute gp_less) moreover have "homotopic_paths (- {z}) h q" by (simp add: e q valid_path_imp_path norm_minus_commute hq_less) ultimately have "homotopic_paths (- {z}) p q" by (blast intro: homotopic_paths_trans homotopic_paths_sym assms) then have "contour_integral p (\w. 1/(w - z)) = contour_integral q (\w. 1/(w - z))" by (rule Cauchy_theorem_homotopic_paths) (auto intro!: holomorphic_intros simp: p q) then show ?thesis by (simp add: pap paq) qed lemma winding_number_homotopic_loops: assumes "homotopic_loops (-{z}) g h" shows "winding_number g z = winding_number h z" proof - have "path g" "path h" using homotopic_loops_imp_path [OF assms] by auto moreover have pag: "z \ path_image g" and pah: "z \ path_image h" using homotopic_loops_imp_subset [OF assms] by auto moreover have gloop: "pathfinish g = pathstart g" and hloop: "pathfinish h = pathstart h" using homotopic_loops_imp_loop [OF assms] by auto ultimately obtain d e where "d > 0" "e > 0" and d: "\p. \path p; pathfinish p = pathstart p; \t\{0..1}. norm (p t - g t) < d\ \ homotopic_loops (-{z}) g p" and e: "\q. \path q; pathfinish q = pathstart q; \t\{0..1}. norm (q t - h t) < e\ \ homotopic_loops (-{z}) h q" using homotopic_nearby_loops [of g "-{z}"] homotopic_nearby_loops [of h "-{z}"] by force obtain p where p: "valid_path p" "z \ path_image p" "pathstart p = pathstart g" "pathfinish p = pathfinish g" and gp_less:"\t\{0..1}. cmod (g t - p t) < d" and pap: "contour_integral p (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number g z" using winding_number [OF \path g\ pag \0 < d\] unfolding winding_number_prop_def by blast obtain q where q: "valid_path q" "z \ path_image q" "pathstart q = pathstart h" "pathfinish q = pathfinish h" and hq_less: "\t\{0..1}. cmod (h t - q t) < e" and paq: "contour_integral q (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number h z" using winding_number [OF \path h\ pah \0 < e\] unfolding winding_number_prop_def by blast have gp: "homotopic_loops (- {z}) g p" by (simp add: gloop d gp_less norm_minus_commute p valid_path_imp_path) have hq: "homotopic_loops (- {z}) h q" by (simp add: e hloop hq_less norm_minus_commute q valid_path_imp_path) have "contour_integral p (\w. 1/(w - z)) = contour_integral q (\w. 1/(w - z))" proof (rule Cauchy_theorem_homotopic_loops) show "homotopic_loops (- {z}) p q" by (blast intro: homotopic_loops_trans homotopic_loops_sym gp hq assms) qed (auto intro!: holomorphic_intros simp: p q) then show ?thesis by (simp add: pap paq) qed lemma winding_number_paths_linear_eq: "\path g; path h; pathstart h = pathstart g; pathfinish h = pathfinish g; \t. t \ {0..1} \ z \ closed_segment (g t) (h t)\ \ winding_number h z = winding_number g z" by (blast intro: sym homotopic_paths_linear winding_number_homotopic_paths) lemma winding_number_loops_linear_eq: "\path g; path h; pathfinish g = pathstart g; pathfinish h = pathstart h; \t. t \ {0..1} \ z \ closed_segment (g t) (h t)\ \ winding_number h z = winding_number g z" by (blast intro: sym homotopic_loops_linear winding_number_homotopic_loops) lemma winding_number_nearby_paths_eq: "\path g; path h; pathstart h = pathstart g; pathfinish h = pathfinish g; \t. t \ {0..1} \ norm(h t - g t) < norm(g t - z)\ \ winding_number h z = winding_number g z" by (metis segment_bound(2) norm_minus_commute not_le winding_number_paths_linear_eq) lemma winding_number_nearby_loops_eq: "\path g; path h; pathfinish g = pathstart g; pathfinish h = pathstart h; \t. t \ {0..1} \ norm(h t - g t) < norm(g t - z)\ \ winding_number h z = winding_number g z" by (metis segment_bound(2) norm_minus_commute not_le winding_number_loops_linear_eq) lemma winding_number_subpath_combine: assumes "path g" and notin: "z \ path_image g" and "u \ {0..1}" "v \ {0..1}" "w \ {0..1}" shows "winding_number (subpath u v g) z + winding_number (subpath v w g) z = winding_number (subpath u w g) z" (is "?lhs = ?rhs") proof - have "?lhs = winding_number (subpath u v g +++ subpath v w g) z" using assms by (metis (no_types, lifting) path_image_subpath_subset path_subpath pathfinish_subpath pathstart_subpath subsetD winding_number_join) also have "... = ?rhs" by (meson assms homotopic_join_subpaths subset_Compl_singleton winding_number_homotopic_paths) finally show ?thesis . qed text \Winding numbers of circular contours\ proposition winding_number_part_circlepath_pos_less: assumes "s < t" and no: "norm(w - z) < r" shows "0 < Re (winding_number(part_circlepath z r s t) w)" proof - have "0 < r" by (meson no norm_not_less_zero not_le order.strict_trans2) note valid_path_part_circlepath moreover have " w \ path_image (part_circlepath z r s t)" using assms by (auto simp: path_image_def image_def part_circlepath_def norm_mult linepath_def) moreover have "0 < r * (t - s) * (r - cmod (w - z))" using assms by (metis \0 < r\ diff_gt_0_iff_gt mult_pos_pos) ultimately show ?thesis apply (rule winding_number_pos_lt [where e = "r*(t - s)*(r - norm(w - z))"]) apply (simp add: vector_derivative_part_circlepath right_diff_distrib [symmetric] mult_ac mult_le_cancel_left_pos assms \0) using Re_Im_le_cmod [of "w-z" "linepath s t x" for x] by (simp add: exp_Euler cos_of_real sin_of_real part_circlepath_def algebra_simps cos_squared_eq [unfolded power2_eq_square]) qed lemma winding_number_circlepath_centre: "0 < r \ winding_number (circlepath z r) z = 1" apply (rule winding_number_unique_loop) apply (simp_all add: sphere_def valid_path_imp_path) apply (rule_tac x="circlepath z r" in exI) apply (simp add: sphere_def contour_integral_circlepath) done proposition winding_number_circlepath: assumes "norm(w - z) < r" shows "winding_number(circlepath z r) w = 1" proof (cases "w = z") case True then show ?thesis using assms winding_number_circlepath_centre by auto next case False have [simp]: "r > 0" using assms le_less_trans norm_ge_zero by blast define r' where "r' = norm(w - z)" have "r' < r" by (simp add: assms r'_def) have disjo: "cball z r' \ sphere z r = {}" using \r' < r\ by (force simp: cball_def sphere_def) have "winding_number(circlepath z r) w = winding_number(circlepath z r) z" proof (rule winding_number_around_inside [where S = "cball z r'"]) show "winding_number (circlepath z r) z \ 0" by (simp add: winding_number_circlepath_centre) show "cball z r' \ path_image (circlepath z r) = {}" by (simp add: disjo less_eq_real_def) qed (auto simp: r'_def dist_norm norm_minus_commute) also have "\ = 1" by (simp add: winding_number_circlepath_centre) finally show ?thesis . qed lemma no_bounded_connected_component_imp_winding_number_zero: assumes g: "path g" "path_image g \ S" "pathfinish g = pathstart g" "z \ S" and nb: "\z. bounded (connected_component_set (- S) z) \ z \ S" shows "winding_number g z = 0" proof - have "z \ outside (path_image g)" by (metis nb [of z] \path_image g \ S\ \z \ S\ subsetD mem_Collect_eq outside outside_mono) then show ?thesis by (simp add: g winding_number_zero_in_outside) qed lemma no_bounded_path_component_imp_winding_number_zero: assumes g: "path g" "path_image g \ S" "pathfinish g = pathstart g" "z \ S" and nb: "\z. bounded (path_component_set (- S) z) \ z \ S" shows "winding_number g z = 0" by (simp add: bounded_subset nb path_component_subset_connected_component no_bounded_connected_component_imp_winding_number_zero [OF g]) subsection\Winding number for a triangle\ lemma wn_triangle1: assumes "0 \ interior(convex hull {a,b,c})" shows "\ (Im(a/b) \ 0 \ 0 \ Im(b/c))" proof - { assume 0: "Im(a/b) \ 0" "0 \ Im(b/c)" have "0 \ interior (convex hull {a,b,c})" proof (cases "a=0 \ b=0 \ c=0") case True then show ?thesis by (auto simp: not_in_interior_convex_hull_3) next case False then have "b \ 0" by blast { fix x y::complex and u::real assume eq_f': "Im x * Re b \ Im b * Re x" "Im y * Re b \ Im b * Re y" "0 \ u" "u \ 1" then have "((1 - u) * Im x) * Re b \ Im b * ((1 - u) * Re x)" by (simp add: mult_left_mono mult.assoc mult.left_commute [of "Im b"]) then have "((1 - u) * Im x + u * Im y) * Re b \ Im b * ((1 - u) * Re x + u * Re y)" using eq_f' ordered_comm_semiring_class.comm_mult_left_mono by (fastforce simp add: algebra_simps) } then have "convex {z. Im z * Re b \ Im b * Re z}" by (auto simp: algebra_simps convex_alt) with False 0 have "convex hull {a,b,c} \ {z. Im z * Re b \ Im b * Re z}" by (simp add: subset_hull mult.commute Complex.Im_divide divide_simps complex_neq_0 [symmetric]) moreover have "0 \ interior({z. Im z * Re b \ Im b * Re z})" proof assume "0 \ interior {z. Im z * Re b \ Im b * Re z}" then obtain e where "e>0" and e: "ball 0 e \ {z. Im z * Re b \ Im b * Re z}" by (meson mem_interior) define z where "z \ - sgn (Im b) * (e/3) + sgn (Re b) * (e/3) * \" have "cmod z = cmod (e/3) * cmod (- sgn (Im b) + sgn (Re b) * \)" unfolding z_def norm_mult [symmetric] by (simp add: algebra_simps) also have "... < e" using \e>0\ by (auto simp: norm_mult intro: le_less_trans [OF norm_triangle_ineq4]) finally have "z \ ball 0 e" - using \e>0\ by (simp add: ) + using \e>0\ by simp then have "Im z * Re b \ Im b * Re z" using e by blast then have b: "0 < Re b" "0 < Im b" and disj: "e * Re b < - (Im b * e) \ e * Re b = - (Im b * e)" using \e>0\ \b \ 0\ by (auto simp add: z_def dist_norm sgn_if less_eq_real_def mult_less_0_iff complex.expand split: if_split_asm) show False \\or just one smt line\ using disj proof assume "e * Re b < - (Im b * e)" with b \0 < e\ less_trans [of _ 0] show False by (metis (no_types) mult_pos_pos neg_less_0_iff_less not_less_iff_gr_or_eq) next assume "e * Re b = - (Im b * e)" with b \0 < e\ show False by (metis mult_pos_pos neg_less_0_iff_less not_less_iff_gr_or_eq) qed qed ultimately show ?thesis using interior_mono by blast qed } with assms show ?thesis by blast qed lemma wn_triangle2_0: assumes "0 \ interior(convex hull {a,b,c})" shows "0 < Im((b - a) * cnj (b)) \ 0 < Im((c - b) * cnj (c)) \ 0 < Im((a - c) * cnj (a)) \ Im((b - a) * cnj (b)) < 0 \ 0 < Im((b - c) * cnj (b)) \ 0 < Im((a - b) * cnj (a)) \ 0 < Im((c - a) * cnj (c))" proof - have [simp]: "{b,c,a} = {a,b,c}" "{c,a,b} = {a,b,c}" by auto show ?thesis using wn_triangle1 [OF assms] wn_triangle1 [of b c a] wn_triangle1 [of c a b] assms by (auto simp: algebra_simps Im_complex_div_gt_0 Im_complex_div_lt_0 not_le not_less) qed lemma wn_triangle2: assumes "z \ interior(convex hull {a,b,c})" shows "0 < Im((b - a) * cnj (b - z)) \ 0 < Im((c - b) * cnj (c - z)) \ 0 < Im((a - c) * cnj (a - z)) \ Im((b - a) * cnj (b - z)) < 0 \ 0 < Im((b - c) * cnj (b - z)) \ 0 < Im((a - b) * cnj (a - z)) \ 0 < Im((c - a) * cnj (c - z))" proof - have 0: "0 \ interior(convex hull {a-z, b-z, c-z})" using assms convex_hull_translation [of "-z" "{a,b,c}"] interior_translation [of "-z"] by (simp cong: image_cong_simp) show ?thesis using wn_triangle2_0 [OF 0] by simp qed lemma wn_triangle3: assumes z: "z \ interior(convex hull {a,b,c})" and "0 < Im((b-a) * cnj (b-z))" "0 < Im((c-b) * cnj (c-z))" "0 < Im((a-c) * cnj (a-z))" shows "winding_number (linepath a b +++ linepath b c +++ linepath c a) z = 1" proof - have znot[simp]: "z \ closed_segment a b" "z \ closed_segment b c" "z \ closed_segment c a" using z interior_of_triangle [of a b c] by (auto simp: closed_segment_def) have gt0: "0 < Re (winding_number (linepath a b +++ linepath b c +++ linepath c a) z)" using assms by (simp add: winding_number_linepath_pos_lt path_image_join winding_number_join_pos_combined) have lt2: "Re (winding_number (linepath a b +++ linepath b c +++ linepath c a) z) < 2" using winding_number_lt_half_linepath [of _ a b] using winding_number_lt_half_linepath [of _ b c] using winding_number_lt_half_linepath [of _ c a] znot by (fastforce simp add: winding_number_join path_image_join) show ?thesis by (rule winding_number_eq_1) (simp_all add: path_image_join gt0 lt2) qed proposition winding_number_triangle: assumes z: "z \ interior(convex hull {a,b,c})" shows "winding_number(linepath a b +++ linepath b c +++ linepath c a) z = (if 0 < Im((b - a) * cnj (b - z)) then 1 else -1)" proof - have [simp]: "{a,c,b} = {a,b,c}" by auto have znot[simp]: "z \ closed_segment a b" "z \ closed_segment b c" "z \ closed_segment c a" using z interior_of_triangle [of a b c] by (auto simp: closed_segment_def) then have [simp]: "z \ closed_segment b a" "z \ closed_segment c b" "z \ closed_segment a c" using closed_segment_commute by blast+ have *: "winding_number (linepath a b +++ linepath b c +++ linepath c a) z = winding_number (reversepath (linepath a c +++ linepath c b +++ linepath b a)) z" by (simp add: reversepath_joinpaths winding_number_join not_in_path_image_join) show ?thesis apply (rule disjE [OF wn_triangle2 [OF z]]) subgoal by (simp add: wn_triangle3 z) subgoal by (simp add: path_image_join winding_number_reversepath * wn_triangle3 z) done qed subsection\Winding numbers for simple closed paths\ lemma winding_number_from_innerpath: assumes "simple_path c1" and c1: "pathstart c1 = a" "pathfinish c1 = b" and "simple_path c2" and c2: "pathstart c2 = a" "pathfinish c2 = b" and "simple_path c" and c: "pathstart c = a" "pathfinish c = b" and c1c2: "path_image c1 \ path_image c2 = {a,b}" and c1c: "path_image c1 \ path_image c = {a,b}" and c2c: "path_image c2 \ path_image c = {a,b}" and ne_12: "path_image c \ inside(path_image c1 \ path_image c2) \ {}" and z: "z \ inside(path_image c1 \ path_image c)" and wn_d: "winding_number (c1 +++ reversepath c) z = d" and "a \ b" "d \ 0" obtains "z \ inside(path_image c1 \ path_image c2)" "winding_number (c1 +++ reversepath c2) z = d" proof - obtain 0: "inside(path_image c1 \ path_image c) \ inside(path_image c2 \ path_image c) = {}" and 1: "inside(path_image c1 \ path_image c) \ inside(path_image c2 \ path_image c) \ (path_image c - {a,b}) = inside(path_image c1 \ path_image c2)" by (rule split_inside_simple_closed_curve [OF \simple_path c1\ c1 \simple_path c2\ c2 \simple_path c\ c \a \ b\ c1c2 c1c c2c ne_12]) have znot: "z \ path_image c" "z \ path_image c1" "z \ path_image c2" using union_with_outside z 1 by auto then have zout: "z \ outside (path_image c \ path_image c2)" by (metis "0" ComplI UnE disjoint_iff_not_equal sup.commute union_with_inside z) have wn_cc2: "winding_number (c +++ reversepath c2) z = 0" by (rule winding_number_zero_in_outside; simp add: zout \simple_path c2\ c c2 \simple_path c\ simple_path_imp_path path_image_join) show ?thesis proof show "z \ inside (path_image c1 \ path_image c2)" using "1" z by blast have "winding_number c1 z - winding_number c z = d " using assms znot by (metis wn_d winding_number_join simple_path_imp_path winding_number_reversepath add.commute path_image_reversepath path_reversepath pathstart_reversepath uminus_add_conv_diff) then show "winding_number (c1 +++ reversepath c2) z = d" using wn_cc2 by (simp add: winding_number_join simple_path_imp_path assms znot winding_number_reversepath) qed qed lemma simple_closed_path_wn1: fixes a::complex and e::real assumes "0 < e" and sp_pl: "simple_path(p +++ linepath (a - e) (a + e))" (is "simple_path ?pae") and psp: "pathstart p = a + e" and pfp: "pathfinish p = a - e" and disj: "ball a e \ path_image p = {}" obtains z where "z \ inside (path_image ?pae)" "cmod (winding_number ?pae z) = 1" proof - have "arc p" and arc_lp: "arc (linepath (a - e) (a + e))" and pap: "path_image p \ path_image (linepath (a - e) (a + e)) \ {pathstart p, a-e}" using simple_path_join_loop_eq [of "linepath (a - e) (a + e)" p] assms by auto have mid_eq_a: "midpoint (a - e) (a + e) = a" by (simp add: midpoint_def) with assms have "a \ path_image ?pae" by (simp add: assms path_image_join) (metis midpoint_in_closed_segment) then have "a \ frontier(inside (path_image ?pae))" using assms by (simp add: Jordan_inside_outside ) with \0 < e\ obtain c where c: "c \ inside (path_image ?pae)" and dac: "dist a c < e" by (auto simp: frontier_straddle) then have "c \ path_image ?pae" using inside_no_overlap by blast then have "c \ path_image p" "c \ closed_segment (a - e) (a + e)" by (simp_all add: assms path_image_join) with \0 < e\ dac have "c \ affine hull {a - e, a + e}" by (simp add: segment_as_ball not_le) with \0 < e\ have *: "\ collinear {a - e, c,a + e}" using collinear_3_affine_hull [of "a-e" "a+e"] by (auto simp: insert_commute) have 13: "1/3 + 1/3 + 1/3 = (1::real)" by simp have "(1/3) *\<^sub>R (a - of_real e) + (1/3) *\<^sub>R c + (1/3) *\<^sub>R (a + of_real e) \ interior(convex hull {a - e, c, a + e})" using interior_convex_hull_3_minimal [OF * DIM_complex] by clarsimp (metis 13 zero_less_divide_1_iff zero_less_numeral) then obtain z where z: "z \ interior(convex hull {a - e, c, a + e})" by force have [simp]: "z \ closed_segment (a - e) c" by (metis DIM_complex Diff_iff IntD2 inf_sup_absorb interior_of_triangle z) have [simp]: "z \ closed_segment (a + e) (a - e)" by (metis DIM_complex DiffD2 Un_iff interior_of_triangle z) have [simp]: "z \ closed_segment c (a + e)" by (metis (no_types, lifting) DIM_complex DiffD2 Un_insert_right inf_sup_aci(5) insertCI interior_of_triangle mk_disjoint_insert z) show thesis proof have "norm (winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z) = 1" using winding_number_triangle [OF z] by simp have zin: "z \ inside (path_image (linepath (a + e) (a - e)) \ path_image p)" and zeq: "winding_number (linepath (a + e) (a - e) +++ reversepath p) z = winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z" proof (rule winding_number_from_innerpath [of "linepath (a + e) (a - e)" "a+e" "a-e" p "linepath (a + e) c +++ linepath c (a - e)" z "winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z"]) show sp_aec: "simple_path (linepath (a + e) c +++ linepath c (a - e))" proof (rule arc_imp_simple_path [OF arc_join]) show "arc (linepath (a + e) c)" by (metis \c \ path_image p\ arc_linepath pathstart_in_path_image psp) show "arc (linepath c (a - e))" by (metis \c \ path_image p\ arc_linepath pathfinish_in_path_image pfp) show "path_image (linepath (a + e) c) \ path_image (linepath c (a - e)) \ {pathstart (linepath c (a - e))}" by clarsimp (metis "*" IntI Int_closed_segment closed_segment_commute singleton_iff) qed auto show "simple_path p" using \arc p\ arc_simple_path by blast show sp_ae2: "simple_path (linepath (a + e) (a - e))" using \arc p\ arc_distinct_ends pfp psp by fastforce show pa: "pathfinish (linepath (a + e) (a - e)) = a - e" "pathstart (linepath (a + e) c +++ linepath c (a - e)) = a + e" "pathfinish (linepath (a + e) c +++ linepath c (a - e)) = a - e" "pathstart p = a + e" "pathfinish p = a - e" "pathstart (linepath (a + e) (a - e)) = a + e" by (simp_all add: assms) show 1: "path_image (linepath (a + e) (a - e)) \ path_image p = {a + e, a - e}" proof show "path_image (linepath (a + e) (a - e)) \ path_image p \ {a + e, a - e}" using pap closed_segment_commute psp segment_convex_hull by fastforce show "{a + e, a - e} \ path_image (linepath (a + e) (a - e)) \ path_image p" using pap pathfinish_in_path_image pathstart_in_path_image pfp psp by fastforce qed show 2: "path_image (linepath (a + e) (a - e)) \ path_image (linepath (a + e) c +++ linepath c (a - e)) = {a + e, a - e}" (is "?lhs = ?rhs") proof have "\ collinear {c, a + e, a - e}" using * by (simp add: insert_commute) then have "convex hull {a + e, a - e} \ convex hull {a + e, c} = {a + e}" "convex hull {a + e, a - e} \ convex hull {c, a - e} = {a - e}" by (metis (full_types) Int_closed_segment insert_commute segment_convex_hull)+ then show "?lhs \ ?rhs" by (metis Int_Un_distrib equalityD1 insert_is_Un path_image_join path_image_linepath path_join_eq path_linepath segment_convex_hull simple_path_def sp_aec) show "?rhs \ ?lhs" using segment_convex_hull by (simp add: path_image_join) qed have "path_image p \ path_image (linepath (a + e) c) \ {a + e}" proof (clarsimp simp: path_image_join) fix x assume "x \ path_image p" and x_ac: "x \ closed_segment (a + e) c" then have "dist x a \ e" by (metis IntI all_not_in_conv disj dist_commute mem_ball not_less) with x_ac dac \e > 0\ show "x = a + e" by (auto simp: norm_minus_commute dist_norm closed_segment_eq_open dest: open_segment_furthest_le [where y=a]) qed moreover have "path_image p \ path_image (linepath c (a - e)) \ {a - e}" proof (clarsimp simp: path_image_join) fix x assume "x \ path_image p" and x_ac: "x \ closed_segment c (a - e)" then have "dist x a \ e" by (metis IntI all_not_in_conv disj dist_commute mem_ball not_less) with x_ac dac \e > 0\ show "x = a - e" by (auto simp: norm_minus_commute dist_norm closed_segment_eq_open dest: open_segment_furthest_le [where y=a]) qed ultimately have "path_image p \ path_image (linepath (a + e) c +++ linepath c (a - e)) \ {a + e, a - e}" by (force simp: path_image_join) then show 3: "path_image p \ path_image (linepath (a + e) c +++ linepath c (a - e)) = {a + e, a - e}" using "1" "2" by blast show 4: "path_image (linepath (a + e) c +++ linepath c (a - e)) \ inside (path_image (linepath (a + e) (a - e)) \ path_image p) \ {}" apply (clarsimp simp: path_image_join segment_convex_hull disjoint_iff_not_equal) by (metis (no_types, opaque_lifting) UnI1 Un_commute c closed_segment_commute ends_in_segment(1) path_image_join path_image_linepath pathstart_linepath pfp segment_convex_hull) show zin_inside: "z \ inside (path_image (linepath (a + e) (a - e)) \ path_image (linepath (a + e) c +++ linepath c (a - e)))" proof (simp add: path_image_join) show "z \ inside (closed_segment (a + e) (a - e) \ (closed_segment (a + e) c \ closed_segment c (a - e)))" by (metis z inside_of_triangle DIM_complex Un_commute closed_segment_commute) qed show 5: "winding_number (linepath (a + e) (a - e) +++ reversepath (linepath (a + e) c +++ linepath c (a - e))) z = winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z" by (simp add: reversepath_joinpaths path_image_join winding_number_join) show 6: "winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z \ 0" by (simp add: winding_number_triangle z) show "winding_number (linepath (a + e) (a - e) +++ reversepath p) z = winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z" by (metis 1 2 3 4 5 6 pa sp_aec sp_ae2 \arc p\ \simple_path p\ arc_distinct_ends winding_number_from_innerpath zin_inside) qed (use assms \e > 0\ in auto) show z_inside: "z \ inside (path_image ?pae)" using zin by (simp add: assms path_image_join Un_commute closed_segment_commute) have "cmod (winding_number ?pae z) = cmod ((winding_number(reversepath ?pae) z))" proof (subst winding_number_reversepath) show "path ?pae" using simple_path_imp_path sp_pl by blast show "z \ path_image ?pae" by (metis IntI emptyE inside_no_overlap z_inside) qed (simp add: inside_def) also have "... = cmod (winding_number(linepath (a + e) (a - e) +++ reversepath p) z)" by (simp add: pfp reversepath_joinpaths) also have "... = cmod (winding_number (linepath (a - e) c +++ linepath c (a + e) +++ linepath (a + e) (a - e)) z)" by (simp add: zeq) also have "... = 1" using z by (simp add: interior_of_triangle winding_number_triangle) finally show "cmod (winding_number ?pae z) = 1" . qed qed lemma simple_closed_path_wn2: fixes a::complex and d e::real assumes "0 < d" "0 < e" and sp_pl: "simple_path(p +++ linepath (a - d) (a + e))" and psp: "pathstart p = a + e" and pfp: "pathfinish p = a - d" obtains z where "z \ inside (path_image (p +++ linepath (a - d) (a + e)))" "cmod (winding_number (p +++ linepath (a - d) (a + e)) z) = 1" proof - have [simp]: "a + of_real x \ closed_segment (a - \) (a - \) \ x \ closed_segment (-\) (-\)" for x \ \::real using closed_segment_translation_eq [of a] by (metis (no_types, opaque_lifting) add_uminus_conv_diff of_real_minus of_real_closed_segment) have [simp]: "a - of_real x \ closed_segment (a + \) (a + \) \ -x \ closed_segment \ \" for x \ \::real by (metis closed_segment_translation_eq diff_conv_add_uminus of_real_closed_segment of_real_minus) have "arc p" and arc_lp: "arc (linepath (a - d) (a + e))" and "path p" and pap: "path_image p \ closed_segment (a - d) (a + e) \ {a+e, a-d}" using simple_path_join_loop_eq [of "linepath (a - d) (a + e)" p] assms arc_imp_path by auto have "0 \ closed_segment (-d) e" using \0 < d\ \0 < e\ closed_segment_eq_real_ivl by auto then have "a \ path_image (linepath (a - d) (a + e))" using of_real_closed_segment [THEN iffD2] by (force dest: closed_segment_translation_eq [of a, THEN iffD2] simp del: of_real_closed_segment) then have "a \ path_image p" using \0 < d\ \0 < e\ pap by auto then obtain k where "0 < k" and k: "ball a k \ (path_image p) = {}" using \0 < e\ \path p\ not_on_path_ball by blast define kde where "kde \ (min k (min d e)) / 2" have "0 < kde" "kde < k" "kde < d" "kde < e" using \0 < k\ \0 < d\ \0 < e\ by (auto simp: kde_def) let ?q = "linepath (a + kde) (a + e) +++ p +++ linepath (a - d) (a - kde)" have "- kde \ closed_segment (-d) e" using \0 < kde\ \kde < d\ \kde < e\ closed_segment_eq_real_ivl by auto then have a_diff_kde: "a - kde \ closed_segment (a - d) (a + e)" using of_real_closed_segment [THEN iffD2] by (force dest: closed_segment_translation_eq [of a, THEN iffD2] simp del: of_real_closed_segment) then have clsub2: "closed_segment (a - d) (a - kde) \ closed_segment (a - d) (a + e)" by (simp add: subset_closed_segment) then have "path_image p \ closed_segment (a - d) (a - kde) \ {a + e, a - d}" using pap by force moreover have "a + e \ path_image p \ closed_segment (a - d) (a - kde)" using \0 < kde\ \kde < d\ \0 < e\ by (auto simp: closed_segment_eq_real_ivl) ultimately have sub_a_diff_d: "path_image p \ closed_segment (a - d) (a - kde) \ {a - d}" by blast have "kde \ closed_segment (-d) e" using \0 < kde\ \kde < d\ \kde < e\ closed_segment_eq_real_ivl by auto then have a_diff_kde: "a + kde \ closed_segment (a - d) (a + e)" using of_real_closed_segment [THEN iffD2] by (force dest: closed_segment_translation_eq [of "a", THEN iffD2] simp del: of_real_closed_segment) then have clsub1: "closed_segment (a + kde) (a + e) \ closed_segment (a - d) (a + e)" by (simp add: subset_closed_segment) then have "closed_segment (a + kde) (a + e) \ path_image p \ {a + e, a - d}" using pap by force moreover have "closed_segment (a + kde) (a + e) \ closed_segment (a - d) (a - kde) = {}" proof (clarsimp intro!: equals0I) fix y assume y1: "y \ closed_segment (a + kde) (a + e)" and y2: "y \ closed_segment (a - d) (a - kde)" obtain u::real where u: "y = a + u" and "0 < u" proof - obtain \ where \: "y = (1 - \) *\<^sub>R (a + kde) + \ *\<^sub>R (a + e)" and "0 \ \" "\ \ 1" using y1 by (auto simp: in_segment) show thesis proof show "y = a + ((1 - \)*kde + \*e)" using \ by (auto simp: scaleR_conv_of_real algebra_simps) have "(1 - \)*kde + \*e \ 0" using \0 < kde\ \0 \ \\ \\ \ 1\ \0 < e\ by auto moreover have "(1 - \)*kde + \*e \ 0" using \0 < kde\ \0 \ \\ \\ \ 1\ \0 < e\ by (auto simp: add_nonneg_eq_0_iff) ultimately show "(1 - \)*kde + \*e > 0" by simp qed qed moreover obtain v::real where v: "y = a + v" and "v \ 0" proof - obtain \ where \: "y = (1 - \) *\<^sub>R (a - d) + \ *\<^sub>R (a - kde)" and "0 \ \" "\ \ 1" using y2 by (auto simp: in_segment) show thesis proof show "y = a + (- ((1 - \)*d + \*kde))" using \ by (force simp: scaleR_conv_of_real algebra_simps) show "- ((1 - \)*d + \*kde) \ 0" using \0 < kde\ \0 \ \\ \\ \ 1\ \0 < d\ by (metis add.left_neutral add_nonneg_nonneg le_diff_eq less_eq_real_def neg_le_0_iff_le split_mult_pos_le) qed qed ultimately show False by auto qed moreover have "a - d \ closed_segment (a + kde) (a + e)" using \0 < kde\ \kde < d\ \0 < e\ by (auto simp: closed_segment_eq_real_ivl) ultimately have sub_a_plus_e: "closed_segment (a + kde) (a + e) \ (path_image p \ closed_segment (a - d) (a - kde)) \ {a + e}" by auto have "kde \ closed_segment (-kde) e" using \0 < kde\ \kde < d\ \kde < e\ closed_segment_eq_real_ivl by auto then have a_add_kde: "a + kde \ closed_segment (a - kde) (a + e)" using of_real_closed_segment [THEN iffD2] by (force dest: closed_segment_translation_eq [of "a", THEN iffD2] simp del: of_real_closed_segment) have "closed_segment (a - kde) (a + kde) \ closed_segment (a + kde) (a + e) = {a + kde}" by (metis a_add_kde Int_closed_segment) moreover have "path_image p \ closed_segment (a - kde) (a + kde) = {}" proof (rule equals0I, clarify) fix y assume "y \ path_image p" "y \ closed_segment (a - kde) (a + kde)" with equals0D [OF k, of y] \0 < kde\ \kde < k\ show False by (auto simp: dist_norm dest: dist_decreases_closed_segment [where c=a]) qed moreover have "- kde \ closed_segment (-d) kde" using \0 < kde\ \kde < d\ \kde < e\ closed_segment_eq_real_ivl by auto then have a_diff_kde': "a - kde \ closed_segment (a - d) (a + kde)" using of_real_closed_segment [THEN iffD2] by (force dest: closed_segment_translation_eq [of a, THEN iffD2] simp del: of_real_closed_segment) then have "closed_segment (a - d) (a - kde) \ closed_segment (a - kde) (a + kde) = {a - kde}" by (metis Int_closed_segment) ultimately have pa_subset_pm_kde: "path_image ?q \ closed_segment (a - kde) (a + kde) \ {a - kde, a + kde}" by (auto simp: path_image_join assms) have ge_kde1: "\y. x = a + of_real y \ y \ kde" if x: "x \ closed_segment (a + kde) (a + e)" for x proof - obtain u where "0 \ u" "u \ 1" and u: "x = (1 - u) *\<^sub>R (a + kde) + u *\<^sub>R (a + e)" using x by (auto simp: in_segment) then have "kde \ (1 - u) * kde + u * e" using \kde < e\ segment_bound_lemma by auto moreover have "x = a + ((1 - u) * kde + u * e)" by (fastforce simp: u algebra_simps scaleR_conv_of_real) ultimately show ?thesis by blast qed have ge_kde2: "\y. x = a + of_real y \ y \ -kde" if x: "x \ closed_segment (a - d) (a - kde)" for x proof - obtain u where "0 \ u" "u \ 1" and u: "x = (1 - u) *\<^sub>R (a - d) + u *\<^sub>R (a - kde)" using x by (auto simp: in_segment) then have "kde \ ((1-u)*d + u*kde)" using \kde < d\ segment_bound_lemma by auto moreover have "x = a - ((1-u)*d + u*kde)" by (fastforce simp: u algebra_simps scaleR_conv_of_real) ultimately show ?thesis by (metis add_uminus_conv_diff neg_le_iff_le of_real_minus) qed have notin_paq: "x \ path_image ?q" if "dist a x < kde" for x proof - have "x \ path_image p" using k \kde < k\ that by force then show ?thesis using that assms by (auto simp: path_image_join dist_norm dest!: ge_kde1 ge_kde2) qed obtain z where zin: "z \ inside (path_image (?q +++ linepath (a - kde) (a + kde)))" and z1: "cmod (winding_number (?q +++ linepath (a - kde) (a + kde)) z) = 1" proof (rule simple_closed_path_wn1 [of kde ?q a]) show "simple_path (?q +++ linepath (a - kde) (a + kde))" proof (intro simple_path_join_loop conjI) show "arc ?q" proof (rule arc_join) show "arc (linepath (a + kde) (a + e))" using \kde < e\ \arc p\ by (force simp: pfp) show "arc (p +++ linepath (a - d) (a - kde))" using \kde < d\ \kde < e\ \arc p\ sub_a_diff_d by (force simp: pfp intro: arc_join) qed (auto simp: psp pfp path_image_join sub_a_plus_e) show "arc (linepath (a - kde) (a + kde))" using \0 < kde\ by auto qed (use pa_subset_pm_kde in auto) qed (use \0 < kde\ notin_paq in auto) have eq: "path_image (?q +++ linepath (a - kde) (a + kde)) = path_image (p +++ linepath (a - d) (a + e))" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" using clsub1 clsub2 apply (auto simp: path_image_join assms) by (meson subsetCE subset_closed_segment) show "?rhs \ ?lhs" apply (simp add: path_image_join assms Un_ac) by (metis Un_closed_segment Un_assoc a_diff_kde a_diff_kde' le_supI2 subset_refl) qed show thesis proof show zzin: "z \ inside (path_image (p +++ linepath (a - d) (a + e)))" by (metis eq zin) then have znotin: "z \ path_image p" by (metis ComplD Un_iff inside_Un_outside path_image_join pathfinish_linepath pathstart_reversepath pfp reversepath_linepath) have znotin_d_kde: "z \ closed_segment (a - d) (a + kde)" by (metis ComplD Un_iff Un_closed_segment a_diff_kde inside_Un_outside path_image_join path_image_linepath pathstart_linepath pfp zzin) have znotin_d_e: "z \ closed_segment (a - d) (a + e)" by (metis IntI UnCI emptyE inside_no_overlap path_image_join path_image_linepath pathstart_linepath pfp zzin) have znotin_kde_e: "z \ closed_segment (a + kde) (a + e)" and znotin_d_kde': "z \ closed_segment (a - d) (a - kde)" using clsub1 clsub2 zzin by (metis (no_types, opaque_lifting) IntI Un_iff emptyE inside_no_overlap path_image_join path_image_linepath pathstart_linepath pfp subsetD)+ have "winding_number (linepath (a - d) (a + e)) z = winding_number (linepath (a - d) (a + kde)) z + winding_number (linepath (a + kde) (a + e)) z" proof (rule winding_number_split_linepath) show "a + complex_of_real kde \ closed_segment (a - d) (a + e)" by (simp add: a_diff_kde) show "z \ closed_segment (a - d) (a + e)" by (metis ComplD Un_iff inside_Un_outside path_image_join path_image_linepath pathstart_linepath pfp zzin) qed also have "... = winding_number (linepath (a + kde) (a + e)) z + (winding_number (linepath (a - d) (a - kde)) z + winding_number (linepath (a - kde) (a + kde)) z)" by (simp add: winding_number_split_linepath [of "a-kde", symmetric] znotin_d_kde a_diff_kde') finally have "winding_number (p +++ linepath (a - d) (a + e)) z = winding_number p z + winding_number (linepath (a + kde) (a + e)) z + (winding_number (linepath (a - d) (a - kde)) z + winding_number (linepath (a - kde) (a + kde)) z)" by (metis (no_types, lifting) ComplD Un_iff \arc p\ add.assoc arc_imp_path eq path_image_join path_join_path_ends path_linepath simple_path_imp_path sp_pl union_with_outside winding_number_join zin) also have "... = winding_number (linepath (a + kde) (a + e)) z + winding_number (p +++ linepath (a - d) (a - kde)) z + winding_number (linepath (a - kde) (a + kde)) z" using \path p\ znotin assms by simp (metis Un_iff Un_closed_segment a_diff_kde' path_image_linepath path_linepath pathstart_linepath winding_number_join znotin_d_kde) also have "... = winding_number ?q z + winding_number (linepath (a - kde) (a + kde)) z" using \path p\ znotin assms by (simp add: path_image_join winding_number_join znotin_kde_e znotin_d_kde') also have "... = winding_number (?q +++ linepath (a - kde) (a + kde)) z" by (metis (mono_tags, lifting) ComplD UnCI \path p\ eq inside_outside path_image_join path_join_eq path_linepath pathfinish_join pathfinish_linepath pathstart_join pathstart_linepath pfp psp winding_number_join zzin) finally have "winding_number (p +++ linepath (a - d) (a + e)) z = winding_number (?q +++ linepath (a - kde) (a + kde)) z" . then show "cmod (winding_number (p +++ linepath (a - d) (a + e)) z) = 1" by (simp add: z1) qed qed lemma simple_closed_path_wn3: fixes p :: "real \ complex" assumes "simple_path p" and loop: "pathfinish p = pathstart p" obtains z where "z \ inside (path_image p)" "cmod (winding_number p z) = 1" proof - have ins: "inside(path_image p) \ {}" "open(inside(path_image p))" "connected(inside(path_image p))" and out: "outside(path_image p) \ {}" "open(outside(path_image p))" "connected(outside(path_image p))" and bo: "bounded(inside(path_image p))" "\ bounded(outside(path_image p))" and ins_out: "inside(path_image p) \ outside(path_image p) = {}" "inside(path_image p) \ outside(path_image p) = - path_image p" and fro: "frontier(inside(path_image p)) = path_image p" "frontier(outside(path_image p)) = path_image p" using Jordan_inside_outside [OF assms] by auto obtain a where a: "a \ inside(path_image p)" using \inside (path_image p) \ {}\ by blast obtain d::real where "0 < d" and d_fro: "a - d \ frontier (inside (path_image p))" and d_int: "\\. \0 \ \; \ < d\ \ (a - \) \ inside (path_image p)" using ray_to_frontier [of "inside (path_image p)" a "-1"] bo ins a interior_eq by (metis ab_group_add_class.ab_diff_conv_add_uminus of_real_def scale_minus_right zero_neq_neg_one) obtain e::real where "0 < e" and e_fro: "a + e \ frontier (inside (path_image p))" and e_int: "\\. \0 \ \; \ < e\ \ (a + \) \ inside (path_image p)" using ray_to_frontier [of "inside (path_image p)" a 1] bo ins a interior_eq by (metis of_real_def zero_neq_one) obtain t0 where "0 \ t0" "t0 \ 1" and pt: "p t0 = a - d" using a d_fro fro by (auto simp: path_image_def) obtain q where "simple_path q" and q_ends: "pathstart q = a - d" "pathfinish q = a - d" and q_eq_p: "path_image q = path_image p" and wn_q_eq_wn_p: "\z. z \ inside(path_image p) \ winding_number q z = winding_number p z" proof show "simple_path (shiftpath t0 p)" by (simp add: pathstart_shiftpath pathfinish_shiftpath simple_path_shiftpath path_image_shiftpath \0 \ t0\ \t0 \ 1\ assms) show "pathstart (shiftpath t0 p) = a - d" using pt by (simp add: \t0 \ 1\ pathstart_shiftpath) show "pathfinish (shiftpath t0 p) = a - d" by (simp add: \0 \ t0\ loop pathfinish_shiftpath pt) show "path_image (shiftpath t0 p) = path_image p" by (simp add: \0 \ t0\ \t0 \ 1\ loop path_image_shiftpath) show "winding_number (shiftpath t0 p) z = winding_number p z" if "z \ inside (path_image p)" for z by (metis ComplD Un_iff \0 \ t0\ \t0 \ 1\ \simple_path p\ atLeastAtMost_iff inside_Un_outside loop simple_path_imp_path that winding_number_shiftpath) qed have ad_not_ae: "a - d \ a + e" by (metis \0 < d\ \0 < e\ add.left_inverse add_left_cancel add_uminus_conv_diff le_add_same_cancel2 less_eq_real_def not_less of_real_add of_real_def of_real_eq_0_iff pt) have ad_ae_q: "{a - d, a + e} \ path_image q" using \path_image q = path_image p\ d_fro e_fro fro(1) by auto have ada: "open_segment (a - d) a \ inside (path_image p)" proof (clarsimp simp: in_segment) fix u::real assume "0 < u" "u < 1" with d_int have "a - (1 - u) * d \ inside (path_image p)" by (metis \0 < d\ add.commute diff_add_cancel left_diff_distrib' less_add_same_cancel2 less_eq_real_def mult.left_neutral zero_less_mult_iff) then show "(1 - u) *\<^sub>R (a - d) + u *\<^sub>R a \ inside (path_image p)" by (simp add: diff_add_eq of_real_def real_vector.scale_right_diff_distrib) qed have aae: "open_segment a (a + e) \ inside (path_image p)" proof (clarsimp simp: in_segment) fix u::real assume "0 < u" "u < 1" with e_int have "a + u * e \ inside (path_image p)" by (meson \0 < e\ less_eq_real_def mult_less_cancel_right2 not_less zero_less_mult_iff) then show "(1 - u) *\<^sub>R a + u *\<^sub>R (a + e) \ inside (path_image p)" by (metis (mono_tags, lifting) add.assoc of_real_mult pth_6 scaleR_collapse scaleR_conv_of_real) qed have "complex_of_real (d * d + (e * e + d * (e + e))) \ 0" using ad_not_ae by (metis \0 < d\ \0 < e\ add_strict_left_mono less_add_same_cancel1 not_sum_squares_lt_zero of_real_eq_0_iff zero_less_double_add_iff_zero_less_single_add zero_less_mult_iff) moreover have "\u>0. u < 1 \ d = u * d + u * e" using \0 < d\ \0 < e\ by (rule_tac x="d / (d+e)" in exI) (simp add: divide_simps scaleR_conv_of_real flip: distrib_left) ultimately have a_in_de: "a \ open_segment (a - d) (a + e)" using ad_not_ae by (simp add: in_segment algebra_simps scaleR_conv_of_real flip: of_real_add of_real_mult) then have "open_segment (a - d) (a + e) \ inside (path_image p)" using ada a aae Un_open_segment [of a "a-d" "a+e"] by blast then have "path_image q \ open_segment (a - d) (a + e) = {}" using inside_no_overlap by (fastforce simp: q_eq_p) with ad_ae_q have paq_Int_cs: "path_image q \ closed_segment (a - d) (a + e) = {a - d, a + e}" by (simp add: closed_segment_eq_open) obtain t where "0 \ t" "t \ 1" and qt: "q t = a + e" using a e_fro fro ad_ae_q by (auto simp: path_defs) then have "t \ 0" by (metis ad_not_ae pathstart_def q_ends(1)) then have "t \ 1" by (metis ad_not_ae pathfinish_def q_ends(2) qt) have q01: "q 0 = a - d" "q 1 = a - d" using q_ends by (auto simp: pathstart_def pathfinish_def) obtain z where zin: "z \ inside (path_image (subpath t 0 q +++ linepath (a - d) (a + e)))" and z1: "cmod (winding_number (subpath t 0 q +++ linepath (a - d) (a + e)) z) = 1" proof (rule simple_closed_path_wn2 [of d e "subpath t 0 q" a], simp_all add: q01) show "simple_path (subpath t 0 q +++ linepath (a - d) (a + e))" proof (rule simple_path_join_loop, simp_all add: qt q01) have "inj_on q (closed_segment t 0)" using \0 \ t\ \simple_path q\ \t \ 1\ \t \ 0\ \t \ 1\ by (fastforce simp: simple_path_def inj_on_def closed_segment_eq_real_ivl) then show "arc (subpath t 0 q)" using \0 \ t\ \simple_path q\ \t \ 1\ \t \ 0\ by (simp add: arc_subpath_eq simple_path_imp_path) show "arc (linepath (a - d) (a + e))" by (simp add: ad_not_ae) show "path_image (subpath t 0 q) \ closed_segment (a - d) (a + e) \ {a + e, a - d}" using qt paq_Int_cs \simple_path q\ \0 \ t\ \t \ 1\ by (force simp: dest: rev_subsetD [OF _ path_image_subpath_subset] intro: simple_path_imp_path) qed qed (auto simp: \0 < d\ \0 < e\ qt) have pa01_Un: "path_image (subpath 0 t q) \ path_image (subpath 1 t q) = path_image q" unfolding path_image_subpath using \0 \ t\ \t \ 1\ by (force simp: path_image_def image_iff) with paq_Int_cs have pa_01q: "(path_image (subpath 0 t q) \ path_image (subpath 1 t q)) \ closed_segment (a - d) (a + e) = {a - d, a + e}" by metis have z_notin_ed: "z \ closed_segment (a + e) (a - d)" using zin q01 by (simp add: path_image_join closed_segment_commute inside_def) have z_notin_0t: "z \ path_image (subpath 0 t q)" by (metis (no_types, opaque_lifting) IntI Un_upper1 subsetD empty_iff inside_no_overlap path_image_join path_image_subpath_commute pathfinish_subpath pathstart_def pathstart_linepath q_ends(1) qt subpath_trivial zin) have [simp]: "- winding_number (subpath t 0 q) z = winding_number (subpath 0 t q) z" by (metis \0 \ t\ \simple_path q\ \t \ 1\ atLeastAtMost_iff zero_le_one path_image_subpath_commute path_subpath real_eq_0_iff_le_ge_0 reversepath_subpath simple_path_imp_path winding_number_reversepath z_notin_0t) obtain z_in_q: "z \ inside(path_image q)" and wn_q: "winding_number (subpath 0 t q +++ subpath t 1 q) z = - winding_number (subpath t 0 q +++ linepath (a - d) (a + e)) z" proof (rule winding_number_from_innerpath [of "subpath 0 t q" "a-d" "a+e" "subpath 1 t q" "linepath (a - d) (a + e)" z "- winding_number (subpath t 0 q +++ linepath (a - d) (a + e)) z"], simp_all add: q01 qt pa01_Un reversepath_subpath) show "simple_path (subpath 0 t q)" "simple_path (subpath 1 t q)" by (simp_all add: \0 \ t\ \simple_path q\ \t \ 1\ \t \ 0\ \t \ 1\ simple_path_subpath) show "simple_path (linepath (a - d) (a + e))" using ad_not_ae by blast show "path_image (subpath 0 t q) \ path_image (subpath 1 t q) = {a - d, a + e}" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" using \0 \ t\ \simple_path q\ \t \ 1\ \t \ 1\ q_ends qt q01 by (force simp: pathfinish_def qt simple_path_def path_image_subpath) show "?rhs \ ?lhs" using \0 \ t\ \t \ 1\ q01 qt by (force simp: path_image_subpath) qed show "path_image (subpath 0 t q) \ closed_segment (a - d) (a + e) = {a - d, a + e}" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" using paq_Int_cs pa01_Un by fastforce show "?rhs \ ?lhs" using \0 \ t\ \t \ 1\ q01 qt by (force simp: path_image_subpath) qed show "path_image (subpath 1 t q) \ closed_segment (a - d) (a + e) = {a - d, a + e}" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (auto simp: pa_01q [symmetric]) show "?rhs \ ?lhs" using \0 \ t\ \t \ 1\ q01 qt by (force simp: path_image_subpath) qed show "closed_segment (a - d) (a + e) \ inside (path_image q) \ {}" using a a_in_de open_closed_segment pa01_Un q_eq_p by fastforce show "z \ inside (path_image (subpath 0 t q) \ closed_segment (a - d) (a + e))" by (metis path_image_join path_image_linepath path_image_subpath_commute pathfinish_subpath pathstart_linepath q01(1) zin) show "winding_number (subpath 0 t q +++ linepath (a + e) (a - d)) z = - winding_number (subpath t 0 q +++ linepath (a - d) (a + e)) z" using z_notin_ed z_notin_0t \0 \ t\ \simple_path q\ \t \ 1\ by (simp add: simple_path_imp_path qt q01 path_image_subpath_commute closed_segment_commute winding_number_join winding_number_reversepath [symmetric]) show "- d \ e" using ad_not_ae by auto show "winding_number (subpath t 0 q +++ linepath (a - d) (a + e)) z \ 0" using z1 by auto qed show ?thesis proof show "z \ inside (path_image p)" using q_eq_p z_in_q by auto then have [simp]: "z \ path_image q" by (metis disjoint_iff_not_equal inside_no_overlap q_eq_p) have [simp]: "z \ path_image (subpath 1 t q)" using inside_def pa01_Un z_in_q by fastforce have "winding_number(subpath 0 t q +++ subpath t 1 q) z = winding_number(subpath 0 1 q) z" using z_notin_0t \0 \ t\ \simple_path q\ \t \ 1\ by (simp add: simple_path_imp_path qt path_image_subpath_commute winding_number_join winding_number_subpath_combine) with wn_q have "winding_number (subpath t 0 q +++ linepath (a - d) (a + e)) z = - winding_number q z" by auto with z1 have "cmod (winding_number q z) = 1" by simp with z1 wn_q_eq_wn_p show "cmod (winding_number p z) = 1" using z1 wn_q_eq_wn_p by (simp add: \z \ inside (path_image p)\) qed qed proposition simple_closed_path_winding_number_inside: assumes "simple_path \" obtains "\z. z \ inside(path_image \) \ winding_number \ z = 1" | "\z. z \ inside(path_image \) \ winding_number \ z = -1" proof (cases "pathfinish \ = pathstart \") case True have "path \" by (simp add: assms simple_path_imp_path) then have const: "winding_number \ constant_on inside(path_image \)" proof (rule winding_number_constant) show "connected (inside(path_image \))" by (simp add: Jordan_inside_outside True assms) qed (use inside_no_overlap True in auto) obtain z where zin: "z \ inside (path_image \)" and z1: "cmod (winding_number \ z) = 1" using simple_closed_path_wn3 [of \] True assms by blast have "winding_number \ z \ \" using zin integer_winding_number [OF \path \\ True] inside_def by blast moreover have "real_of_int x = - 1 \ x = -1" for x by linarith ultimately consider "winding_number \ z = 1" | "winding_number \ z = -1" using z1 by (auto simp: Ints_def abs_if split: if_split_asm) with that const zin show ?thesis unfolding constant_on_def by metis next case False then show ?thesis using inside_simple_curve_imp_closed assms that(2) by blast qed lemma simple_closed_path_abs_winding_number_inside: assumes "simple_path \" "z \ inside(path_image \)" shows "\Re (winding_number \ z)\ = 1" by (metis assms norm_minus_cancel norm_one one_complex.simps(1) real_norm_def simple_closed_path_winding_number_inside uminus_complex.simps(1)) lemma simple_closed_path_norm_winding_number_inside: assumes "simple_path \" "z \ inside(path_image \)" shows "norm (winding_number \ z) = 1" proof - have "pathfinish \ = pathstart \" using assms inside_simple_curve_imp_closed by blast with assms integer_winding_number have "winding_number \ z \ \" by (simp add: inside_def simple_path_def) then show ?thesis by (metis assms norm_minus_cancel norm_one simple_closed_path_winding_number_inside) qed lemma simple_closed_path_winding_number_cases: assumes "simple_path \" "pathfinish \ = pathstart \" "z \ path_image \" shows "winding_number \ z \ {-1,0,1}" proof - consider "z \ inside (path_image \)" | "z \ outside (path_image \)" by (metis ComplI UnE assms(3) inside_Un_outside) then show ?thesis proof cases case 1 then show ?thesis using assms simple_closed_path_winding_number_inside by auto next case 2 then show ?thesis using assms simple_path_def winding_number_zero_in_outside by blast qed qed lemma simple_closed_path_winding_number_pos: "\simple_path \; pathfinish \ = pathstart \; z \ path_image \; 0 < Re(winding_number \ z)\ \ winding_number \ z = 1" using simple_closed_path_winding_number_cases by fastforce subsection \Winding number for rectangular paths\ proposition winding_number_rectpath: assumes "z \ box a1 a3" shows "winding_number (rectpath a1 a3) z = 1" proof - from assms have less: "Re a1 < Re a3" "Im a1 < Im a3" by (auto simp: in_box_complex_iff) define a2 a4 where "a2 = Complex (Re a3) (Im a1)" and "a4 = Complex (Re a1) (Im a3)" let ?l1 = "linepath a1 a2" and ?l2 = "linepath a2 a3" and ?l3 = "linepath a3 a4" and ?l4 = "linepath a4 a1" from assms and less have "z \ path_image (rectpath a1 a3)" by (auto simp: path_image_rectpath_cbox_minus_box) also have "path_image (rectpath a1 a3) = path_image ?l1 \ path_image ?l2 \ path_image ?l3 \ path_image ?l4" by (simp add: rectpath_def Let_def path_image_join Un_assoc a2_def a4_def) finally have "z \ \" . moreover have "\l\{?l1,?l2,?l3,?l4}. Re (winding_number l z) > 0" unfolding ball_simps HOL.simp_thms a2_def a4_def by (intro conjI; (rule winding_number_linepath_pos_lt; (insert assms, auto simp: a2_def a4_def in_box_complex_iff mult_neg_neg))+) ultimately have "Re (winding_number (rectpath a1 a3) z) > 0" by (simp add: winding_number_join path_image_join rectpath_def Let_def a2_def a4_def) thus "winding_number (rectpath a1 a3) z = 1" using assms less by (intro simple_closed_path_winding_number_pos simple_path_rectpath) (auto simp: path_image_rectpath_cbox_minus_box) qed proposition winding_number_rectpath_outside: assumes "Re a1 \ Re a3" "Im a1 \ Im a3" assumes "z \ cbox a1 a3" shows "winding_number (rectpath a1 a3) z = 0" using assms by (intro winding_number_zero_outside[OF _ _ _ assms(3)] path_image_rectpath_subset_cbox) simp_all text\A per-function version for continuous logs, a kind of monodromy\ proposition\<^marker>\tag unimportant\ winding_number_compose_exp: assumes "path p" shows "winding_number (exp \ p) 0 = (pathfinish p - pathstart p) / (2 * of_real pi * \)" proof - obtain e where "0 < e" and e: "\t. t \ {0..1} \ e \ norm(exp(p t))" proof have "closed (path_image (exp \ p))" by (simp add: assms closed_path_image holomorphic_on_exp holomorphic_on_imp_continuous_on path_continuous_image) then show "0 < setdist {0} (path_image (exp \ p))" by (metis exp_not_eq_zero imageE image_comp infdist_eq_setdist infdist_pos_not_in_closed path_defs(4) path_image_nonempty) next fix t::real assume "t \ {0..1}" have "setdist {0} (path_image (exp \ p)) \ dist 0 (exp (p t))" proof (rule setdist_le_dist) show "exp (p t) \ path_image (exp \ p)" using \t \ {0..1}\ path_image_def by fastforce+ qed auto then show "setdist {0} (path_image (exp \ p)) \ cmod (exp (p t))" by simp qed have "bounded (path_image p)" by (simp add: assms bounded_path_image) then obtain B where "0 < B" and B: "path_image p \ cball 0 B" by (meson bounded_pos mem_cball_0 subsetI) let ?B = "cball (0::complex) (B+1)" have "uniformly_continuous_on ?B exp" using holomorphic_on_exp holomorphic_on_imp_continuous_on by (force intro: compact_uniformly_continuous) then obtain d where "d > 0" and d: "\x x'. \x\?B; x'\?B; dist x' x < d\ \ norm (exp x' - exp x) < e" using \e > 0\ by (auto simp: uniformly_continuous_on_def dist_norm) then have "min 1 d > 0" by force then obtain g where pfg: "polynomial_function g" and "g 0 = p 0" "g 1 = p 1" and gless: "\t. t \ {0..1} \ norm(g t - p t) < min 1 d" using path_approx_polynomial_function [OF \path p\] \d > 0\ \0 < e\ unfolding pathfinish_def pathstart_def by blast have "winding_number (exp \ p) 0 = winding_number (exp \ g) 0" proof (rule winding_number_nearby_paths_eq [symmetric]) show "path (exp \ p)" "path (exp \ g)" by (simp_all add: pfg assms holomorphic_on_exp holomorphic_on_imp_continuous_on path_continuous_image path_polynomial_function) next fix t :: "real" assume t: "t \ {0..1}" with gless have "norm(g t - p t) < 1" using min_less_iff_conj by blast moreover have ptB: "norm (p t) \ B" using B t by (force simp: path_image_def) ultimately have "cmod (g t) \ B + 1" by (meson add_mono_thms_linordered_field(4) le_less_trans less_imp_le norm_triangle_sub) with ptB gless t have "cmod ((exp \ g) t - (exp \ p) t) < e" by (auto simp: dist_norm d) with e t show "cmod ((exp \ g) t - (exp \ p) t) < cmod ((exp \ p) t - 0)" by fastforce qed (use \g 0 = p 0\ \g 1 = p 1\ in \auto simp: pathfinish_def pathstart_def\) also have "... = 1 / (of_real (2 * pi) * \) * contour_integral (exp \ g) (\w. 1 / (w - 0))" proof (rule winding_number_valid_path) have "continuous_on (path_image g) (deriv exp)" by (metis DERIV_exp DERIV_imp_deriv continuous_on_cong holomorphic_on_exp holomorphic_on_imp_continuous_on) then show "valid_path (exp \ g)" by (simp add: field_differentiable_within_exp pfg valid_path_compose valid_path_polynomial_function) show "0 \ path_image (exp \ g)" by (auto simp: path_image_def) qed also have "... = 1 / (of_real (2 * pi) * \) * integral {0..1} (\x. vector_derivative g (at x))" proof (simp add: contour_integral_integral, rule integral_cong) fix t :: "real" assume t: "t \ {0..1}" show "vector_derivative (exp \ g) (at t) / exp (g t) = vector_derivative g (at t)" proof - have "(exp \ g has_vector_derivative vector_derivative (exp \ g) (at t)) (at t)" by (meson DERIV_exp differentiable_def field_vector_diff_chain_at has_vector_derivative_def has_vector_derivative_polynomial_function pfg vector_derivative_works) moreover have "(exp \ g has_vector_derivative vector_derivative g (at t) * exp (g t)) (at t)" by (metis DERIV_exp field_vector_diff_chain_at has_vector_derivative_polynomial_function pfg vector_derivative_at) ultimately show ?thesis by (simp add: divide_simps, rule vector_derivative_unique_at) qed qed also have "... = (pathfinish p - pathstart p) / (2 * of_real pi * \)" proof - have "((\x. vector_derivative g (at x)) has_integral g 1 - g 0) {0..1}" by (meson differentiable_at_polynomial_function fundamental_theorem_of_calculus has_vector_derivative_at_within pfg vector_derivative_works zero_le_one) then show ?thesis unfolding pathfinish_def pathstart_def using \g 0 = p 0\ \g 1 = p 1\ by auto qed finally show ?thesis . qed end \ No newline at end of file diff --git a/src/HOL/Computational_Algebra/Nth_Powers.thy b/src/HOL/Computational_Algebra/Nth_Powers.thy --- a/src/HOL/Computational_Algebra/Nth_Powers.thy +++ b/src/HOL/Computational_Algebra/Nth_Powers.thy @@ -1,308 +1,308 @@ (* File: HOL/Computational_Algebra/Nth_Powers.thy Author: Manuel Eberl n-th powers in general and n-th roots of natural numbers *) section \$n$-th powers and roots of naturals\ theory Nth_Powers imports Primes begin subsection \The set of $n$-th powers\ definition is_nth_power :: "nat \ 'a :: monoid_mult \ bool" where "is_nth_power n x \ (\y. x = y ^ n)" lemma is_nth_power_nth_power [simp, intro]: "is_nth_power n (x ^ n)" by (auto simp add: is_nth_power_def) lemma is_nth_powerI [intro?]: "x = y ^ n \ is_nth_power n x" by (auto simp: is_nth_power_def) lemma is_nth_powerE: "is_nth_power n x \ (\y. x = y ^ n \ P) \ P" by (auto simp: is_nth_power_def) abbreviation is_square where "is_square \ is_nth_power 2" lemma is_zeroth_power [simp]: "is_nth_power 0 x \ x = 1" by (simp add: is_nth_power_def) lemma is_first_power [simp]: "is_nth_power 1 x" by (simp add: is_nth_power_def) lemma is_first_power' [simp]: "is_nth_power (Suc 0) x" by (simp add: is_nth_power_def) lemma is_nth_power_0 [simp]: "n > 0 \ is_nth_power n (0 :: 'a :: semiring_1)" by (auto simp: is_nth_power_def power_0_left intro!: exI[of _ 0]) lemma is_nth_power_0_iff [simp]: "is_nth_power n (0 :: 'a :: semiring_1) \ n > 0" by (cases n) auto lemma is_nth_power_1 [simp]: "is_nth_power n 1" by (auto simp: is_nth_power_def intro!: exI[of _ 1]) lemma is_nth_power_Suc_0 [simp]: "is_nth_power n (Suc 0)" by (simp add: One_nat_def [symmetric] del: One_nat_def) lemma is_nth_power_conv_multiplicity: fixes x :: "'a :: {factorial_semiring, normalization_semidom_multiplicative}" assumes "n > 0" shows "is_nth_power n (normalize x) \ (\p. prime p \ n dvd multiplicity p x)" proof (cases "x = 0") case False show ?thesis proof (safe intro!: is_nth_powerI elim!: is_nth_powerE) fix y p :: 'a assume *: "normalize x = y ^ n" "prime p" with assms and False have [simp]: "y \ 0" by (auto simp: power_0_left) have "multiplicity p x = multiplicity p (y ^ n)" by (subst *(1) [symmetric]) simp with False and * and assms show "n dvd multiplicity p x" by (auto simp: prime_elem_multiplicity_power_distrib) next assume *: "\p. prime p \ n dvd multiplicity p x" have "multiplicity p ((\p\prime_factors x. p ^ (multiplicity p x div n)) ^ n) = multiplicity p x" if "prime p" for p proof - from that and * have "n dvd multiplicity p x" by blast have "multiplicity p x = 0" if "p \ prime_factors x" using that and \prime p\ by (simp add: prime_factors_multiplicity) with that and * and assms show ?thesis unfolding prod_power_distrib power_mult [symmetric] by (subst multiplicity_prod_prime_powers) (auto simp: in_prime_factors_imp_prime elim: dvdE) qed with assms False have "normalize x = normalize ((\p\prime_factors x. p ^ (multiplicity p x div n)) ^ n)" by (intro multiplicity_eq_imp_eq) (auto simp: multiplicity_prod_prime_powers) thus "normalize x = normalize (\p\prime_factors x. p ^ (multiplicity p x div n)) ^ n" by (simp add: normalize_power) qed qed (insert assms, auto) lemma is_nth_power_conv_multiplicity_nat: assumes "n > 0" shows "is_nth_power n (x :: nat) \ (\p. prime p \ n dvd multiplicity p x)" using is_nth_power_conv_multiplicity[OF assms, of x] by simp lemma is_nth_power_mult: assumes "is_nth_power n a" "is_nth_power n b" shows "is_nth_power n (a * b :: 'a :: comm_monoid_mult)" proof - from assms obtain a' b' where "a = a' ^ n" "b = b' ^ n" by (auto elim!: is_nth_powerE) hence "a * b = (a' * b') ^ n" by (simp add: power_mult_distrib) thus ?thesis by (rule is_nth_powerI) qed lemma is_nth_power_mult_coprime_natD: fixes a b :: nat assumes "coprime a b" "is_nth_power n (a * b)" "a > 0" "b > 0" shows "is_nth_power n a" "is_nth_power n b" proof - have A: "is_nth_power n a" if "coprime a b" "is_nth_power n (a * b)" "a \ 0" "b \ 0" "n > 0" for a b :: nat unfolding is_nth_power_conv_multiplicity_nat[OF \n > 0\] proof safe fix p :: nat assume p: "prime p" from \coprime a b\ have "\(p dvd a \ p dvd b)" using coprime_common_divisor_nat[of a b p] p by auto moreover from that and p have "n dvd multiplicity p a + multiplicity p b" by (auto simp: is_nth_power_conv_multiplicity_nat prime_elem_multiplicity_mult_distrib) ultimately show "n dvd multiplicity p a" by (auto simp: not_dvd_imp_multiplicity_0) qed from A [of a b] assms show "is_nth_power n a" by (cases "n = 0") simp_all from A [of b a] assms show "is_nth_power n b" by (cases "n = 0") (simp_all add: ac_simps) qed lemma is_nth_power_mult_coprime_nat_iff: fixes a b :: nat assumes "coprime a b" shows "is_nth_power n (a * b) \ is_nth_power n a \is_nth_power n b" using assms by (cases "a = 0"; cases "b = 0") (auto intro: is_nth_power_mult dest: is_nth_power_mult_coprime_natD[of a b n] simp del: One_nat_def) lemma is_nth_power_prime_power_nat_iff: fixes p :: nat assumes "prime p" shows "is_nth_power n (p ^ k) \ n dvd k" using assms by (cases "n > 0") (auto simp: is_nth_power_conv_multiplicity_nat prime_elem_multiplicity_power_distrib) lemma is_nth_power_nth_power': assumes "n dvd n'" shows "is_nth_power n (m ^ n')" proof - from assms have "n' = n' div n * n" by simp also have "m ^ \ = (m ^ (n' div n)) ^ n" by (simp add: power_mult) also have "is_nth_power n \" by simp finally show ?thesis . qed definition is_nth_power_nat :: "nat \ nat \ bool" where [code_abbrev]: "is_nth_power_nat = is_nth_power" lemma is_nth_power_nat_code [code]: "is_nth_power_nat n m = (if n = 0 then m = 1 else if m = 0 then n > 0 else if n = 1 then True else (\k\{1..m}. k ^ n = m))" by (auto simp: is_nth_power_nat_def is_nth_power_def power_eq_iff_eq_base self_le_power) (* TODO: Harmonise with Discrete.sqrt *) subsection \The $n$-root of a natural number\ definition nth_root_nat :: "nat \ nat \ nat" where "nth_root_nat k n = (if k = 0 then 0 else Max {m. m ^ k \ n})" lemma zeroth_root_nat [simp]: "nth_root_nat 0 n = 0" by (simp add: nth_root_nat_def) lemma nth_root_nat_aux1: assumes "k > 0" shows "{m::nat. m ^ k \ n} \ {..n}" proof safe fix m assume "m ^ k \ n" show "m \ n" proof (cases "m = 0") case False with assms have "m ^ 1 \ m ^ k" by (intro power_increasing) simp_all also note \m ^ k \ n\ finally show ?thesis by simp qed simp_all qed lemma nth_root_nat_aux2: assumes "k > 0" shows "finite {m::nat. m ^ k \ n}" "{m::nat. m ^ k \ n} \ {}" proof - from assms have "{m. m ^ k \ n} \ {..n}" by (rule nth_root_nat_aux1) moreover have "finite {..n}" by simp ultimately show "finite {m::nat. m ^ k \ n}" by (rule finite_subset) next from assms show "{m::nat. m ^ k \ n} \ {}" by (auto intro!: exI[of _ 0] simp: power_0_left) qed lemma assumes "k > 0" shows nth_root_nat_power_le: "nth_root_nat k n ^ k \ n" and nth_root_nat_ge: "x ^ k \ n \ x \ nth_root_nat k n" using Max_in[OF nth_root_nat_aux2[OF assms], of n] Max_ge[OF nth_root_nat_aux2(1)[OF assms], of x n] assms by (auto simp: nth_root_nat_def) lemma nth_root_nat_less: assumes "k > 0" "x ^ k > n" shows "nth_root_nat k n < x" proof - from \k > 0\ have "nth_root_nat k n ^ k \ n" by (rule nth_root_nat_power_le) also have "n < x ^ k" by fact finally show ?thesis by (rule power_less_imp_less_base) simp_all qed lemma nth_root_nat_unique: assumes "m ^ k \ n" "(m + 1) ^ k > n" shows "nth_root_nat k n = m" proof (cases "k > 0") case True from nth_root_nat_less[OF \k > 0\ assms(2)] have "nth_root_nat k n \ m" by simp moreover from \k > 0\ and assms(1) have "nth_root_nat k n \ m" by (intro nth_root_nat_ge) ultimately show ?thesis by (rule antisym) qed (insert assms, auto) lemma nth_root_nat_0 [simp]: "nth_root_nat k 0 = 0" by (simp add: nth_root_nat_def) lemma nth_root_nat_1 [simp]: "k > 0 \ nth_root_nat k 1 = 1" by (rule nth_root_nat_unique) (auto simp del: One_nat_def) lemma nth_root_nat_Suc_0 [simp]: "k > 0 \ nth_root_nat k (Suc 0) = Suc 0" using nth_root_nat_1 by (simp del: nth_root_nat_1) lemma first_root_nat [simp]: "nth_root_nat 1 n = n" by (intro nth_root_nat_unique) auto lemma first_root_nat' [simp]: "nth_root_nat (Suc 0) n = n" by (intro nth_root_nat_unique) auto lemma nth_root_nat_code_naive': "nth_root_nat k n = (if k = 0 then 0 else Max (Set.filter (\m. m ^ k \ n) {..n}))" proof (cases "k > 0") case True hence "{m. m ^ k \ n} \ {..n}" by (rule nth_root_nat_aux1) hence "Set.filter (\m. m ^ k \ n) {..n} = {m. m ^ k \ n}" by (auto simp: Set.filter_def) with True show ?thesis by (simp add: nth_root_nat_def Set.filter_def) qed simp function nth_root_nat_aux :: "nat \ nat \ nat \ nat \ nat" where "nth_root_nat_aux m k acc n = (let acc' = (k + 1) ^ m in if k \ n \ acc' > n then k else nth_root_nat_aux m (k+1) acc' n)" by auto termination by (relation "measure (\(_,k,_,n). n - k)", goal_cases) auto lemma nth_root_nat_aux_le: assumes "k ^ m \ n" "m > 0" shows "nth_root_nat_aux m k (k ^ m) n ^ m \ n" using assms by (induction m k "k ^ m" n rule: nth_root_nat_aux.induct) (auto simp: Let_def) lemma nth_root_nat_aux_gt: assumes "m > 0" shows "(nth_root_nat_aux m k (k ^ m) n + 1) ^ m > n" using assms proof (induction m k "k ^ m" n rule: nth_root_nat_aux.induct) case (1 m k n) have "n < Suc k ^ m" if "n \ k" proof - note that also have "k < Suc k ^ 1" by simp also from \m > 0\ have "\ \ Suc k ^ m" by (intro power_increasing) simp_all finally show ?thesis . qed with 1 show ?case by (auto simp: Let_def) qed lemma nth_root_nat_aux_correct: assumes "k ^ m \ n" "m > 0" shows "nth_root_nat_aux m k (k ^ m) n = nth_root_nat m n" by (rule sym, intro nth_root_nat_unique nth_root_nat_aux_le nth_root_nat_aux_gt assms) lemma nth_root_nat_naive_code [code]: "nth_root_nat m n = (if m = 0 \ n = 0 then 0 else if m = 1 \ n = 1 then n else nth_root_nat_aux m 1 1 n)" - using nth_root_nat_aux_correct[of 1 m n] by (auto simp: ) + using nth_root_nat_aux_correct[of 1 m n] by auto lemma nth_root_nat_nth_power [simp]: "k > 0 \ nth_root_nat k (n ^ k) = n" by (intro nth_root_nat_unique order.refl power_strict_mono) simp_all lemma nth_root_nat_nth_power': assumes "k > 0" "k dvd m" shows "nth_root_nat k (n ^ m) = n ^ (m div k)" proof - from assms have "m = (m div k) * k" by simp also have "n ^ \ = (n ^ (m div k)) ^ k" by (simp add: power_mult) also from assms have "nth_root_nat k \ = n ^ (m div k)" by simp finally show ?thesis . qed lemma nth_root_nat_mono: assumes "m \ n" shows "nth_root_nat k m \ nth_root_nat k n" proof (cases "k = 0") case False with assms show ?thesis unfolding nth_root_nat_def using nth_root_nat_aux2[of k m] nth_root_nat_aux2[of k n] by (auto intro!: Max_mono) qed auto end diff --git a/src/HOL/Decision_Procs/Approximation_Bounds.thy b/src/HOL/Decision_Procs/Approximation_Bounds.thy --- a/src/HOL/Decision_Procs/Approximation_Bounds.thy +++ b/src/HOL/Decision_Procs/Approximation_Bounds.thy @@ -1,3101 +1,3101 @@ (* Author: Johannes Hoelzl, TU Muenchen Coercions removed by Dmitriy Traytel This file contains only general material about computing lower/upper bounds on real functions. Approximation.thy contains the actual approximation algorithm and the approximation oracle. This is in order to make a clear separation between "morally immaculate" material about upper/lower bounds and the trusted oracle/reflection. *) theory Approximation_Bounds imports Complex_Main "HOL-Library.Interval_Float" Dense_Linear_Order begin declare powr_neg_one [simp] declare powr_neg_numeral [simp] context includes interval.lifting begin section "Horner Scheme" subsection \Define auxiliary helper \horner\ function\ primrec horner :: "(nat \ nat) \ (nat \ nat \ nat) \ nat \ nat \ nat \ real \ real" where "horner F G 0 i k x = 0" | "horner F G (Suc n) i k x = 1 / k - x * horner F G n (F i) (G i k) x" lemma horner_schema': fixes x :: real and a :: "nat \ real" shows "a 0 - x * (\ i=0.. i=0..i. - (x * ((-1)^i * a (Suc i) * x ^ i)) = (-1)^(Suc i) * a (Suc i) * x ^ (Suc i)" by auto show ?thesis unfolding sum_distrib_left shift_pow uminus_add_conv_diff [symmetric] sum_negf[symmetric] sum.atLeast_Suc_lessThan[OF zero_less_Suc] sum.reindex[OF inj_Suc, unfolded comp_def, symmetric, of "\ n. (-1)^n *a n * x^n"] by auto qed lemma horner_schema: fixes f :: "nat \ nat" and G :: "nat \ nat \ nat" and F :: "nat \ nat" assumes f_Suc: "\n. f (Suc n) = G ((F ^^ n) s) (f n)" shows "horner F G n ((F ^^ j') s) (f j') x = (\ j = 0..< n. (- 1) ^ j * (1 / (f (j' + j))) * x ^ j)" proof (induct n arbitrary: j') case 0 then show ?case by auto next case (Suc n) show ?case unfolding horner.simps Suc[where j'="Suc j'", unfolded funpow.simps comp_def f_Suc] using horner_schema'[of "\ j. 1 / (f (j' + j))"] by auto qed lemma horner_bounds': fixes lb :: "nat \ nat \ nat \ float \ float" and ub :: "nat \ nat \ nat \ float \ float" assumes "0 \ real_of_float x" and f_Suc: "\n. f (Suc n) = G ((F ^^ n) s) (f n)" and lb_0: "\ i k x. lb 0 i k x = 0" and lb_Suc: "\ n i k x. lb (Suc n) i k x = float_plus_down prec (lapprox_rat prec 1 k) (- float_round_up prec (x * (ub n (F i) (G i k) x)))" and ub_0: "\ i k x. ub 0 i k x = 0" and ub_Suc: "\ n i k x. ub (Suc n) i k x = float_plus_up prec (rapprox_rat prec 1 k) (- float_round_down prec (x * (lb n (F i) (G i k) x)))" shows "(lb n ((F ^^ j') s) (f j') x) \ horner F G n ((F ^^ j') s) (f j') x \ horner F G n ((F ^^ j') s) (f j') x \ (ub n ((F ^^ j') s) (f j') x)" (is "?lb n j' \ ?horner n j' \ ?horner n j' \ ?ub n j'") proof (induct n arbitrary: j') case 0 thus ?case unfolding lb_0 ub_0 horner.simps by auto next case (Suc n) thus ?case using lapprox_rat[of prec 1 "f j'"] using rapprox_rat[of 1 "f j'" prec] Suc[where j'="Suc j'"] \0 \ real_of_float x\ by (auto intro!: add_mono mult_left_mono float_round_down_le float_round_up_le order_trans[OF add_mono[OF _ float_plus_down_le]] order_trans[OF _ add_mono[OF _ float_plus_up_le]] simp add: lb_Suc ub_Suc field_simps f_Suc) qed subsection "Theorems for floating point functions implementing the horner scheme" text \ Here \<^term_type>\f :: nat \ nat\ is the sequence defining the Taylor series, the coefficients are all alternating and reciprocs. We use \<^term>\G\ and \<^term>\F\ to describe the computation of \<^term>\f\. \ lemma horner_bounds: fixes F :: "nat \ nat" and G :: "nat \ nat \ nat" assumes "0 \ real_of_float x" and f_Suc: "\n. f (Suc n) = G ((F ^^ n) s) (f n)" and lb_0: "\ i k x. lb 0 i k x = 0" and lb_Suc: "\ n i k x. lb (Suc n) i k x = float_plus_down prec (lapprox_rat prec 1 k) (- float_round_up prec (x * (ub n (F i) (G i k) x)))" and ub_0: "\ i k x. ub 0 i k x = 0" and ub_Suc: "\ n i k x. ub (Suc n) i k x = float_plus_up prec (rapprox_rat prec 1 k) (- float_round_down prec (x * (lb n (F i) (G i k) x)))" shows "(lb n ((F ^^ j') s) (f j') x) \ (\j=0..j=0.. (ub n ((F ^^ j') s) (f j') x)" (is "?ub") proof - have "?lb \ ?ub" using horner_bounds'[where lb=lb, OF \0 \ real_of_float x\ f_Suc lb_0 lb_Suc ub_0 ub_Suc] unfolding horner_schema[where f=f, OF f_Suc] by simp thus "?lb" and "?ub" by auto qed lemma horner_bounds_nonpos: fixes F :: "nat \ nat" and G :: "nat \ nat \ nat" assumes "real_of_float x \ 0" and f_Suc: "\n. f (Suc n) = G ((F ^^ n) s) (f n)" and lb_0: "\ i k x. lb 0 i k x = 0" and lb_Suc: "\ n i k x. lb (Suc n) i k x = float_plus_down prec (lapprox_rat prec 1 k) (float_round_down prec (x * (ub n (F i) (G i k) x)))" and ub_0: "\ i k x. ub 0 i k x = 0" and ub_Suc: "\ n i k x. ub (Suc n) i k x = float_plus_up prec (rapprox_rat prec 1 k) (float_round_up prec (x * (lb n (F i) (G i k) x)))" shows "(lb n ((F ^^ j') s) (f j') x) \ (\j=0..j=0.. (ub n ((F ^^ j') s) (f j') x)" (is "?ub") proof - have diff_mult_minus: "x - y * z = x + - y * z" for x y z :: float by simp have sum_eq: "(\j=0..j = 0.. real_of_float (-x)" using assms by auto from horner_bounds[where G=G and F=F and f=f and s=s and prec=prec and lb="\ n i k x. lb n i k (-x)" and ub="\ n i k x. ub n i k (-x)", unfolded lb_Suc ub_Suc diff_mult_minus, OF this f_Suc lb_0 _ ub_0 _] show "?lb" and "?ub" unfolding minus_minus sum_eq by (auto simp: minus_float_round_up_eq minus_float_round_down_eq) qed subsection \Selectors for next even or odd number\ text \ The horner scheme computes alternating series. To get the upper and lower bounds we need to guarantee to access a even or odd member. To do this we use \<^term>\get_odd\ and \<^term>\get_even\. \ definition get_odd :: "nat \ nat" where "get_odd n = (if odd n then n else (Suc n))" definition get_even :: "nat \ nat" where "get_even n = (if even n then n else (Suc n))" lemma get_odd[simp]: "odd (get_odd n)" unfolding get_odd_def by (cases "odd n") auto lemma get_even[simp]: "even (get_even n)" unfolding get_even_def by (cases "even n") auto lemma get_odd_ex: "\ k. Suc k = get_odd n \ odd (Suc k)" by (auto simp: get_odd_def odd_pos intro!: exI[of _ "n - 1"]) lemma get_even_double: "\i. get_even n = 2 * i" using get_even by (blast elim: evenE) lemma get_odd_double: "\i. get_odd n = 2 * i + 1" using get_odd by (blast elim: oddE) section "Power function" definition float_power_bnds :: "nat \ nat \ float \ float \ float * float" where "float_power_bnds prec n l u = (if 0 < l then (power_down_fl prec l n, power_up_fl prec u n) else if odd n then (- power_up_fl prec \l\ n, if u < 0 then - power_down_fl prec \u\ n else power_up_fl prec u n) else if u < 0 then (power_down_fl prec \u\ n, power_up_fl prec \l\ n) else (0, power_up_fl prec (max \l\ \u\) n))" lemma le_minus_power_downI: "0 \ x \ x ^ n \ - a \ a \ - power_down prec x n" by (subst le_minus_iff) (auto intro: power_down_le power_mono_odd) lemma float_power_bnds: "(l1, u1) = float_power_bnds prec n l u \ x \ {l .. u} \ (x::real) ^ n \ {l1..u1}" by (auto simp: float_power_bnds_def max_def real_power_up_fl real_power_down_fl minus_le_iff split: if_split_asm intro!: power_up_le power_down_le le_minus_power_downI intro: power_mono_odd power_mono power_mono_even zero_le_even_power) lemma bnds_power: "\(x::real) l u. (l1, u1) = float_power_bnds prec n l u \ x \ {l .. u} \ l1 \ x ^ n \ x ^ n \ u1" using float_power_bnds by auto lift_definition power_float_interval :: "nat \ nat \ float interval \ float interval" is "\p n (l, u). float_power_bnds p n l u" using float_power_bnds by (auto simp: bnds_power dest!: float_power_bnds[OF sym]) lemma lower_power_float_interval: "lower (power_float_interval p n x) = fst (float_power_bnds p n (lower x) (upper x))" by transfer auto lemma upper_power_float_interval: "upper (power_float_interval p n x) = snd (float_power_bnds p n (lower x) (upper x))" by transfer auto lemma power_float_intervalI: "x \\<^sub>r X \ x ^ n \\<^sub>r power_float_interval p n X" using float_power_bnds[OF prod.collapse] by (auto simp: set_of_eq lower_power_float_interval upper_power_float_interval) lemma power_float_interval_mono: "set_of (power_float_interval prec n A) \ set_of (power_float_interval prec n B)" if "set_of A \ set_of B" proof - define la where "la = real_of_float (lower A)" define ua where "ua = real_of_float (upper A)" define lb where "lb = real_of_float (lower B)" define ub where "ub = real_of_float (upper B)" have ineqs: "lb \ la" "la \ ua" "ua \ ub" "lb \ ub" using that lower_le_upper[of A] lower_le_upper[of B] by (auto simp: la_def ua_def lb_def ub_def set_of_eq) show ?thesis using ineqs by (simp add: set_of_subset_iff float_power_bnds_def max_def power_down_fl.rep_eq power_up_fl.rep_eq lower_power_float_interval upper_power_float_interval la_def[symmetric] ua_def[symmetric] lb_def[symmetric] ub_def[symmetric]) (auto intro!: power_down_mono power_up_mono intro: order_trans[where y=0]) qed section \Approximation utility functions\ lift_definition plus_float_interval::"nat \ float interval \ float interval \ float interval" is "\prec. \(a1, a2). \(b1, b2). (float_plus_down prec a1 b1, float_plus_up prec a2 b2)" by (auto intro!: add_mono simp: float_plus_down_le float_plus_up_le) lemma lower_plus_float_interval: "lower (plus_float_interval prec ivl ivl') = float_plus_down prec (lower ivl) (lower ivl')" by transfer auto lemma upper_plus_float_interval: "upper (plus_float_interval prec ivl ivl') = float_plus_up prec (upper ivl) (upper ivl')" by transfer auto lemma mult_float_interval_ge: "real_interval A + real_interval B \ real_interval (plus_float_interval prec A B)" unfolding less_eq_interval_def by transfer (auto simp: lower_plus_float_interval upper_plus_float_interval intro!: order.trans[OF float_plus_down] order.trans[OF _ float_plus_up]) lemma plus_float_interval: "set_of (real_interval A) + set_of (real_interval B) \ set_of (real_interval (plus_float_interval prec A B))" proof - have "set_of (real_interval A) + set_of (real_interval B) \ set_of (real_interval A + real_interval B)" by (simp add: set_of_plus) also have "\ \ set_of (real_interval (plus_float_interval prec A B))" using mult_float_interval_ge[of A B prec] by (simp add: set_of_subset_iff') finally show ?thesis . qed lemma plus_float_intervalI: "x + y \\<^sub>r plus_float_interval prec A B" if "x \\<^sub>i real_interval A" "y \\<^sub>i real_interval B" using plus_float_interval[of A B] that by auto lemma plus_float_interval_mono: "plus_float_interval prec A B \ plus_float_interval prec X Y" if "A \ X" "B \ Y" using that by (auto simp: less_eq_interval_def lower_plus_float_interval upper_plus_float_interval float_plus_down.rep_eq float_plus_up.rep_eq plus_down_mono plus_up_mono) lemma plus_float_interval_monotonic: "set_of (ivl + ivl') \ set_of (plus_float_interval prec ivl ivl')" using float_plus_down_le float_plus_up_le lower_plus_float_interval upper_plus_float_interval by (simp add: set_of_subset_iff) definition bnds_mult :: "nat \ float \ float \ float \ float \ float \ float" where "bnds_mult prec a1 a2 b1 b2 = (float_plus_down prec (nprt a1 * pprt b2) (float_plus_down prec (nprt a2 * nprt b2) (float_plus_down prec (pprt a1 * pprt b1) (pprt a2 * nprt b1))), float_plus_up prec (pprt a2 * pprt b2) (float_plus_up prec (pprt a1 * nprt b2) (float_plus_up prec (nprt a2 * pprt b1) (nprt a1 * nprt b1))))" lemma bnds_mult: fixes prec :: nat and a1 aa2 b1 b2 :: float assumes "(l, u) = bnds_mult prec a1 a2 b1 b2" assumes "a \ {real_of_float a1..real_of_float a2}" assumes "b \ {real_of_float b1..real_of_float b2}" shows "a * b \ {real_of_float l..real_of_float u}" proof - from assms have "real_of_float l \ a * b" by (intro order.trans[OF _ mult_ge_prts[of a1 a a2 b1 b b2]]) (auto simp: bnds_mult_def intro!: float_plus_down_le) moreover from assms have "real_of_float u \ a * b" by (intro order.trans[OF mult_le_prts[of a1 a a2 b1 b b2]]) (auto simp: bnds_mult_def intro!: float_plus_up_le) ultimately show ?thesis by simp qed lift_definition mult_float_interval::"nat \ float interval \ float interval \ float interval" is "\prec. \(a1, a2). \(b1, b2). bnds_mult prec a1 a2 b1 b2" by (auto dest!: bnds_mult[OF sym]) lemma lower_mult_float_interval: "lower (mult_float_interval p x y) = fst (bnds_mult p (lower x) (upper x) (lower y) (upper y))" by transfer auto lemma upper_mult_float_interval: "upper (mult_float_interval p x y) = snd (bnds_mult p (lower x) (upper x) (lower y) (upper y))" by transfer auto lemma mult_float_interval: "set_of (real_interval A) * set_of (real_interval B) \ set_of (real_interval (mult_float_interval prec A B))" proof - let ?bm = "bnds_mult prec (lower A) (upper A) (lower B) (upper B)" show ?thesis using bnds_mult[of "fst ?bm" "snd ?bm", simplified, OF refl] by (auto simp: set_of_eq set_times_def upper_mult_float_interval lower_mult_float_interval) qed lemma mult_float_intervalI: "x * y \\<^sub>r mult_float_interval prec A B" if "x \\<^sub>i real_interval A" "y \\<^sub>i real_interval B" using mult_float_interval[of A B] that - by (auto simp: ) + by auto lemma mult_float_interval_mono': "set_of (mult_float_interval prec A B) \ set_of (mult_float_interval prec X Y)" if "set_of A \ set_of X" "set_of B \ set_of Y" using that apply transfer unfolding bnds_mult_def atLeastatMost_subset_iff float_plus_down.rep_eq float_plus_up.rep_eq by (auto simp: float_plus_down.rep_eq float_plus_up.rep_eq mult_float_mono1 mult_float_mono2) lemma mult_float_interval_mono: "mult_float_interval prec A B \ mult_float_interval prec X Y" if "A \ X" "B \ Y" using mult_float_interval_mono'[of A X B Y prec] that by (simp add: set_of_subset_iff') definition map_bnds :: "(nat \ float \ float) \ (nat \ float \ float) \ nat \ (float \ float) \ (float \ float)" where "map_bnds lb ub prec = (\(l,u). (lb prec l, ub prec u))" lemma map_bnds: assumes "(lf, uf) = map_bnds lb ub prec (l, u)" assumes "mono f" assumes "x \ {real_of_float l..real_of_float u}" assumes "real_of_float (lb prec l) \ f (real_of_float l)" assumes "real_of_float (ub prec u) \ f (real_of_float u)" shows "f x \ {real_of_float lf..real_of_float uf}" proof - from assms have "real_of_float lf = real_of_float (lb prec l)" by (simp add: map_bnds_def) also have "real_of_float (lb prec l) \ f (real_of_float l)" by fact also from assms have "\ \ f x" by (intro monoD[OF \mono f\]) auto finally have lf: "real_of_float lf \ f x" . from assms have "f x \ f (real_of_float u)" by (intro monoD[OF \mono f\]) auto also have "\ \ real_of_float (ub prec u)" by fact also from assms have "\ = real_of_float uf" by (simp add: map_bnds_def) finally have uf: "f x \ real_of_float uf" . from lf uf show ?thesis by simp qed section "Square root" text \ The square root computation is implemented as newton iteration. As first first step we use the nearest power of two greater than the square root. \ fun sqrt_iteration :: "nat \ nat \ float \ float" where "sqrt_iteration prec 0 x = Float 1 ((bitlen \mantissa x\ + exponent x) div 2 + 1)" | "sqrt_iteration prec (Suc m) x = (let y = sqrt_iteration prec m x in Float 1 (- 1) * float_plus_up prec y (float_divr prec x y))" lemma compute_sqrt_iteration_base[code]: shows "sqrt_iteration prec n (Float m e) = (if n = 0 then Float 1 ((if m = 0 then 0 else bitlen \m\ + e) div 2 + 1) else (let y = sqrt_iteration prec (n - 1) (Float m e) in Float 1 (- 1) * float_plus_up prec y (float_divr prec (Float m e) y)))" using bitlen_Float by (cases n) simp_all function ub_sqrt lb_sqrt :: "nat \ float \ float" where "ub_sqrt prec x = (if 0 < x then (sqrt_iteration prec prec x) else if x < 0 then - lb_sqrt prec (- x) else 0)" | "lb_sqrt prec x = (if 0 < x then (float_divl prec x (sqrt_iteration prec prec x)) else if x < 0 then - ub_sqrt prec (- x) else 0)" by pat_completeness auto termination by (relation "measure (\ v. let (prec, x) = case_sum id id v in (if x < 0 then 1 else 0))", auto) declare lb_sqrt.simps[simp del] declare ub_sqrt.simps[simp del] lemma sqrt_ub_pos_pos_1: assumes "sqrt x < b" and "0 < b" and "0 < x" shows "sqrt x < (b + x / b)/2" proof - from assms have "0 < (b - sqrt x)\<^sup>2 " by simp also have "\ = b\<^sup>2 - 2 * b * sqrt x + (sqrt x)\<^sup>2" by algebra also have "\ = b\<^sup>2 - 2 * b * sqrt x + x" using assms by simp finally have "0 < b\<^sup>2 - 2 * b * sqrt x + x" . hence "0 < b / 2 - sqrt x + x / (2 * b)" using assms by (simp add: field_simps power2_eq_square) thus ?thesis by (simp add: field_simps) qed lemma sqrt_iteration_bound: assumes "0 < real_of_float x" shows "sqrt x < sqrt_iteration prec n x" proof (induct n) case 0 show ?case proof (cases x) case (Float m e) hence "0 < m" using assms by (auto simp: algebra_split_simps) hence "0 < sqrt m" by auto have int_nat_bl: "(nat (bitlen m)) = bitlen m" using bitlen_nonneg by auto have "x = (m / 2^nat (bitlen m)) * 2 powr (e + (nat (bitlen m)))" unfolding Float by (auto simp: powr_realpow[symmetric] field_simps powr_add) also have "\ < 1 * 2 powr (e + nat (bitlen m))" proof (rule mult_strict_right_mono, auto) show "m < 2^nat (bitlen m)" using bitlen_bounds[OF \0 < m\, THEN conjunct2] unfolding of_int_less_iff[of m, symmetric] by auto qed finally have "sqrt x < sqrt (2 powr (e + bitlen m))" unfolding int_nat_bl by auto also have "\ \ 2 powr ((e + bitlen m) div 2 + 1)" proof - let ?E = "e + bitlen m" have E_mod_pow: "2 powr (?E mod 2) < 4" proof (cases "?E mod 2 = 1") case True thus ?thesis by auto next case False have "0 \ ?E mod 2" by auto have "?E mod 2 < 2" by auto from this[THEN zless_imp_add1_zle] have "?E mod 2 \ 0" using False by auto from xt1(5)[OF \0 \ ?E mod 2\ this] show ?thesis by auto qed hence "sqrt (2 powr (?E mod 2)) < sqrt (2 * 2)" by (intro real_sqrt_less_mono) auto hence E_mod_pow: "sqrt (2 powr (?E mod 2)) < 2" by auto have E_eq: "2 powr ?E = 2 powr (?E div 2 + ?E div 2 + ?E mod 2)" by auto have "sqrt (2 powr ?E) = sqrt (2 powr (?E div 2) * 2 powr (?E div 2) * 2 powr (?E mod 2))" unfolding E_eq unfolding powr_add[symmetric] by (metis of_int_add) also have "\ = 2 powr (?E div 2) * sqrt (2 powr (?E mod 2))" unfolding real_sqrt_mult[of _ "2 powr (?E mod 2)"] real_sqrt_abs2 by auto also have "\ < 2 powr (?E div 2) * 2 powr 1" by (rule mult_strict_left_mono) (auto intro: E_mod_pow) also have "\ = 2 powr (?E div 2 + 1)" unfolding add.commute[of _ 1] powr_add[symmetric] by simp finally show ?thesis by auto qed finally show ?thesis using \0 < m\ unfolding Float by (subst compute_sqrt_iteration_base) (simp add: ac_simps) qed next case (Suc n) let ?b = "sqrt_iteration prec n x" have "0 < sqrt x" using \0 < real_of_float x\ by auto also have "\ < real_of_float ?b" using Suc . finally have "sqrt x < (?b + x / ?b)/2" using sqrt_ub_pos_pos_1[OF Suc _ \0 < real_of_float x\] by auto also have "\ \ (?b + (float_divr prec x ?b))/2" by (rule divide_right_mono, auto simp add: float_divr) also have "\ = (Float 1 (- 1)) * (?b + (float_divr prec x ?b))" by simp also have "\ \ (Float 1 (- 1)) * (float_plus_up prec ?b (float_divr prec x ?b))" by (auto simp add: algebra_simps float_plus_up_le) finally show ?case unfolding sqrt_iteration.simps Let_def distrib_left . qed lemma sqrt_iteration_lower_bound: assumes "0 < real_of_float x" shows "0 < real_of_float (sqrt_iteration prec n x)" (is "0 < ?sqrt") proof - have "0 < sqrt x" using assms by auto also have "\ < ?sqrt" using sqrt_iteration_bound[OF assms] . finally show ?thesis . qed lemma lb_sqrt_lower_bound: assumes "0 \ real_of_float x" shows "0 \ real_of_float (lb_sqrt prec x)" proof (cases "0 < x") case True hence "0 < real_of_float x" and "0 \ x" using \0 \ real_of_float x\ by auto hence "0 < sqrt_iteration prec prec x" using sqrt_iteration_lower_bound by auto hence "0 \ real_of_float (float_divl prec x (sqrt_iteration prec prec x))" using float_divl_lower_bound[OF \0 \ x\] unfolding less_eq_float_def by auto thus ?thesis unfolding lb_sqrt.simps using True by auto next case False with \0 \ real_of_float x\ have "real_of_float x = 0" by auto thus ?thesis unfolding lb_sqrt.simps by auto qed lemma bnds_sqrt': "sqrt x \ {(lb_sqrt prec x) .. (ub_sqrt prec x)}" proof - have lb: "lb_sqrt prec x \ sqrt x" if "0 < x" for x :: float proof - from that have "0 < real_of_float x" and "0 \ real_of_float x" by auto hence sqrt_gt0: "0 < sqrt x" by auto hence sqrt_ub: "sqrt x < sqrt_iteration prec prec x" using sqrt_iteration_bound by auto have "(float_divl prec x (sqrt_iteration prec prec x)) \ x / (sqrt_iteration prec prec x)" by (rule float_divl) also have "\ < x / sqrt x" by (rule divide_strict_left_mono[OF sqrt_ub \0 < real_of_float x\ mult_pos_pos[OF order_less_trans[OF sqrt_gt0 sqrt_ub] sqrt_gt0]]) also have "\ = sqrt x" unfolding inverse_eq_iff_eq[of _ "sqrt x", symmetric] sqrt_divide_self_eq[OF \0 \ real_of_float x\, symmetric] by auto finally show ?thesis unfolding lb_sqrt.simps if_P[OF \0 < x\] by auto qed have ub: "sqrt x \ ub_sqrt prec x" if "0 < x" for x :: float proof - from that have "0 < real_of_float x" by auto hence "0 < sqrt x" by auto hence "sqrt x < sqrt_iteration prec prec x" using sqrt_iteration_bound by auto then show ?thesis unfolding ub_sqrt.simps if_P[OF \0 < x\] by auto qed show ?thesis using lb[of "-x"] ub[of "-x"] lb[of x] ub[of x] by (auto simp add: lb_sqrt.simps ub_sqrt.simps real_sqrt_minus) qed lemma bnds_sqrt: "\(x::real) lx ux. (l, u) = (lb_sqrt prec lx, ub_sqrt prec ux) \ x \ {lx .. ux} \ l \ sqrt x \ sqrt x \ u" proof ((rule allI) +, rule impI, erule conjE, rule conjI) fix x :: real fix lx ux assume "(l, u) = (lb_sqrt prec lx, ub_sqrt prec ux)" and x: "x \ {lx .. ux}" hence l: "l = lb_sqrt prec lx " and u: "u = ub_sqrt prec ux" by auto have "sqrt lx \ sqrt x" using x by auto from order_trans[OF _ this] show "l \ sqrt x" unfolding l using bnds_sqrt'[of lx prec] by auto have "sqrt x \ sqrt ux" using x by auto from order_trans[OF this] show "sqrt x \ u" unfolding u using bnds_sqrt'[of ux prec] by auto qed lift_definition sqrt_float_interval::"nat \ float interval \ float interval" is "\prec. \(lx, ux). (lb_sqrt prec lx, ub_sqrt prec ux)" using bnds_sqrt' by auto (meson order_trans real_sqrt_le_iff) lemma lower_float_interval: "lower (sqrt_float_interval prec X) = lb_sqrt prec (lower X)" by transfer auto lemma upper_float_interval: "upper (sqrt_float_interval prec X) = ub_sqrt prec (upper X)" by transfer auto lemma sqrt_float_interval: "sqrt ` set_of (real_interval X) \ set_of (real_interval (sqrt_float_interval prec X))" using bnds_sqrt by (auto simp: set_of_eq lower_float_interval upper_float_interval) lemma sqrt_float_intervalI: "sqrt x \\<^sub>r sqrt_float_interval p X" if "x \\<^sub>r X" using sqrt_float_interval[of X p] that by auto section "Arcus tangens and \" subsection "Compute arcus tangens series" text \ As first step we implement the computation of the arcus tangens series. This is only valid in the range \<^term>\{-1 :: real .. 1}\. This is used to compute \ and then the entire arcus tangens. \ fun ub_arctan_horner :: "nat \ nat \ nat \ float \ float" and lb_arctan_horner :: "nat \ nat \ nat \ float \ float" where "ub_arctan_horner prec 0 k x = 0" | "ub_arctan_horner prec (Suc n) k x = float_plus_up prec (rapprox_rat prec 1 k) (- float_round_down prec (x * (lb_arctan_horner prec n (k + 2) x)))" | "lb_arctan_horner prec 0 k x = 0" | "lb_arctan_horner prec (Suc n) k x = float_plus_down prec (lapprox_rat prec 1 k) (- float_round_up prec (x * (ub_arctan_horner prec n (k + 2) x)))" lemma arctan_0_1_bounds': assumes "0 \ real_of_float y" "real_of_float y \ 1" and "even n" shows "arctan (sqrt y) \ {(sqrt y * lb_arctan_horner prec n 1 y) .. (sqrt y * ub_arctan_horner prec (Suc n) 1 y)}" proof - let ?c = "\i. (- 1) ^ i * (1 / (i * 2 + (1::nat)) * sqrt y ^ (i * 2 + 1))" let ?S = "\n. \ i=0.. sqrt y" using assms by auto have "sqrt y \ 1" using assms by auto from \even n\ obtain m where "2 * m = n" by (blast elim: evenE) have "arctan (sqrt y) \ { ?S n .. ?S (Suc n) }" proof (cases "sqrt y = 0") case True then show ?thesis by simp next case False hence "0 < sqrt y" using \0 \ sqrt y\ by auto hence prem: "0 < 1 / (0 * 2 + (1::nat)) * sqrt y ^ (0 * 2 + 1)" by auto have "\ sqrt y \ \ 1" using \0 \ sqrt y\ \sqrt y \ 1\ by auto from mp[OF summable_Leibniz(2)[OF zeroseq_arctan_series[OF this] monoseq_arctan_series[OF this]] prem, THEN spec, of m, unfolded \2 * m = n\] show ?thesis unfolding arctan_series[OF \\ sqrt y \ \ 1\] Suc_eq_plus1 atLeast0LessThan . qed note arctan_bounds = this[unfolded atLeastAtMost_iff] have F: "\n. 2 * Suc n + 1 = 2 * n + 1 + 2" by auto note bounds = horner_bounds[where s=1 and f="\i. 2 * i + 1" and j'=0 and lb="\n i k x. lb_arctan_horner prec n k x" and ub="\n i k x. ub_arctan_horner prec n k x", OF \0 \ real_of_float y\ F lb_arctan_horner.simps ub_arctan_horner.simps] have "(sqrt y * lb_arctan_horner prec n 1 y) \ arctan (sqrt y)" proof - have "(sqrt y * lb_arctan_horner prec n 1 y) \ ?S n" using bounds(1) \0 \ sqrt y\ apply (simp only: power_add power_one_right mult.assoc[symmetric] sum_distrib_right[symmetric]) apply (simp only: mult.commute[where 'a=real] mult.commute[of _ "2::nat"] power_mult) apply (auto intro!: mult_left_mono) done also have "\ \ arctan (sqrt y)" using arctan_bounds .. finally show ?thesis . qed moreover have "arctan (sqrt y) \ (sqrt y * ub_arctan_horner prec (Suc n) 1 y)" proof - have "arctan (sqrt y) \ ?S (Suc n)" using arctan_bounds .. also have "\ \ (sqrt y * ub_arctan_horner prec (Suc n) 1 y)" using bounds(2)[of "Suc n"] \0 \ sqrt y\ apply (simp only: power_add power_one_right mult.assoc[symmetric] sum_distrib_right[symmetric]) apply (simp only: mult.commute[where 'a=real] mult.commute[of _ "2::nat"] power_mult) apply (auto intro!: mult_left_mono) done finally show ?thesis . qed ultimately show ?thesis by auto qed lemma arctan_0_1_bounds: assumes "0 \ real_of_float y" "real_of_float y \ 1" shows "arctan (sqrt y) \ {(sqrt y * lb_arctan_horner prec (get_even n) 1 y) .. (sqrt y * ub_arctan_horner prec (get_odd n) 1 y)}" using arctan_0_1_bounds'[OF assms, of n prec] arctan_0_1_bounds'[OF assms, of "n + 1" prec] arctan_0_1_bounds'[OF assms, of "n - 1" prec] by (auto simp: get_even_def get_odd_def odd_pos simp del: ub_arctan_horner.simps lb_arctan_horner.simps) lemma arctan_lower_bound: assumes "0 \ x" shows "x / (1 + x\<^sup>2) \ arctan x" (is "?l x \ _") proof - have "?l x - arctan x \ ?l 0 - arctan 0" using assms by (intro DERIV_nonpos_imp_nonincreasing[where f="\x. ?l x - arctan x"]) (auto intro!: derivative_eq_intros simp: add_nonneg_eq_0_iff field_simps) thus ?thesis by simp qed lemma arctan_divide_mono: "0 < x \ x \ y \ arctan y / y \ arctan x / x" by (rule DERIV_nonpos_imp_nonincreasing[where f="\x. arctan x / x"]) (auto intro!: derivative_eq_intros divide_nonpos_nonneg simp: inverse_eq_divide arctan_lower_bound) lemma arctan_mult_mono: "0 \ x \ x \ y \ x * arctan y \ y * arctan x" using arctan_divide_mono[of x y] by (cases "x = 0") (simp_all add: field_simps) lemma arctan_mult_le: assumes "0 \ x" "x \ y" "y * z \ arctan y" shows "x * z \ arctan x" proof (cases "x = 0") case True then show ?thesis by simp next case False with assms have "z \ arctan y / y" by (simp add: field_simps) also have "\ \ arctan x / x" using assms \x \ 0\ by (auto intro!: arctan_divide_mono) finally show ?thesis using assms \x \ 0\ by (simp add: field_simps) qed lemma arctan_le_mult: assumes "0 < x" "x \ y" "arctan x \ x * z" shows "arctan y \ y * z" proof - from assms have "arctan y / y \ arctan x / x" by (auto intro!: arctan_divide_mono) also have "\ \ z" using assms by (auto simp: field_simps) finally show ?thesis using assms by (simp add: field_simps) qed lemma arctan_0_1_bounds_le: assumes "0 \ x" "x \ 1" "0 < real_of_float xl" "real_of_float xl \ x * x" "x * x \ real_of_float xu" "real_of_float xu \ 1" shows "arctan x \ {x * lb_arctan_horner p1 (get_even n) 1 xu .. x * ub_arctan_horner p2 (get_odd n) 1 xl}" proof - from assms have "real_of_float xl \ 1" "sqrt (real_of_float xl) \ x" "x \ sqrt (real_of_float xu)" "0 \ real_of_float xu" "0 \ real_of_float xl" "0 < sqrt (real_of_float xl)" by (auto intro!: real_le_rsqrt real_le_lsqrt simp: power2_eq_square) from arctan_0_1_bounds[OF \0 \ real_of_float xu\ \real_of_float xu \ 1\] have "sqrt (real_of_float xu) * real_of_float (lb_arctan_horner p1 (get_even n) 1 xu) \ arctan (sqrt (real_of_float xu))" by simp from arctan_mult_le[OF \0 \ x\ \x \ sqrt _\ this] have "x * real_of_float (lb_arctan_horner p1 (get_even n) 1 xu) \ arctan x" . moreover from arctan_0_1_bounds[OF \0 \ real_of_float xl\ \real_of_float xl \ 1\] have "arctan (sqrt (real_of_float xl)) \ sqrt (real_of_float xl) * real_of_float (ub_arctan_horner p2 (get_odd n) 1 xl)" by simp from arctan_le_mult[OF \0 < sqrt xl\ \sqrt xl \ x\ this] have "arctan x \ x * real_of_float (ub_arctan_horner p2 (get_odd n) 1 xl)" . ultimately show ?thesis by simp qed lemma arctan_0_1_bounds_round: assumes "0 \ real_of_float x" "real_of_float x \ 1" shows "arctan x \ {real_of_float x * lb_arctan_horner p1 (get_even n) 1 (float_round_up (Suc p2) (x * x)) .. real_of_float x * ub_arctan_horner p3 (get_odd n) 1 (float_round_down (Suc p4) (x * x))}" using assms apply (cases "x > 0") apply (intro arctan_0_1_bounds_le) apply (auto simp: float_round_down.rep_eq float_round_up.rep_eq intro!: truncate_up_le1 mult_le_one truncate_down_le truncate_up_le truncate_down_pos mult_pos_pos) done subsection "Compute \" definition ub_pi :: "nat \ float" where "ub_pi prec = (let A = rapprox_rat prec 1 5 ; B = lapprox_rat prec 1 239 in ((Float 1 2) * float_plus_up prec ((Float 1 2) * float_round_up prec (A * (ub_arctan_horner prec (get_odd (prec div 4 + 1)) 1 (float_round_down (Suc prec) (A * A))))) (- float_round_down prec (B * (lb_arctan_horner prec (get_even (prec div 14 + 1)) 1 (float_round_up (Suc prec) (B * B)))))))" definition lb_pi :: "nat \ float" where "lb_pi prec = (let A = lapprox_rat prec 1 5 ; B = rapprox_rat prec 1 239 in ((Float 1 2) * float_plus_down prec ((Float 1 2) * float_round_down prec (A * (lb_arctan_horner prec (get_even (prec div 4 + 1)) 1 (float_round_up (Suc prec) (A * A))))) (- float_round_up prec (B * (ub_arctan_horner prec (get_odd (prec div 14 + 1)) 1 (float_round_down (Suc prec) (B * B)))))))" lemma pi_boundaries: "pi \ {(lb_pi n) .. (ub_pi n)}" proof - have machin_pi: "pi = 4 * (4 * arctan (1 / 5) - arctan (1 / 239))" unfolding machin[symmetric] by auto { fix prec n :: nat fix k :: int assume "1 < k" hence "0 \ k" and "0 < k" and "1 \ k" by auto let ?k = "rapprox_rat prec 1 k" let ?kl = "float_round_down (Suc prec) (?k * ?k)" have "1 div k = 0" using div_pos_pos_trivial[OF _ \1 < k\] by auto have "0 \ real_of_float ?k" by (rule order_trans[OF _ rapprox_rat]) (auto simp add: \0 \ k\) have "real_of_float ?k \ 1" by (auto simp add: \0 < k\ \1 \ k\ less_imp_le intro!: mult_le_one order_trans[OF _ rapprox_rat] rapprox_rat_le1) have "1 / k \ ?k" using rapprox_rat[where x=1 and y=k] by auto hence "arctan (1 / k) \ arctan ?k" by (rule arctan_monotone') also have "\ \ (?k * ub_arctan_horner prec (get_odd n) 1 ?kl)" using arctan_0_1_bounds_round[OF \0 \ real_of_float ?k\ \real_of_float ?k \ 1\] by auto finally have "arctan (1 / k) \ ?k * ub_arctan_horner prec (get_odd n) 1 ?kl" . } note ub_arctan = this { fix prec n :: nat fix k :: int assume "1 < k" hence "0 \ k" and "0 < k" by auto let ?k = "lapprox_rat prec 1 k" let ?ku = "float_round_up (Suc prec) (?k * ?k)" have "1 div k = 0" using div_pos_pos_trivial[OF _ \1 < k\] by auto have "1 / k \ 1" using \1 < k\ by auto have "0 \ real_of_float ?k" using lapprox_rat_nonneg[where x=1 and y=k, OF zero_le_one \0 \ k\] by (auto simp add: \1 div k = 0\) have "0 \ real_of_float (?k * ?k)" by simp have "real_of_float ?k \ 1" using lapprox_rat by (rule order_trans, auto simp add: \1 / k \ 1\) hence "real_of_float (?k * ?k) \ 1" using \0 \ real_of_float ?k\ by (auto intro!: mult_le_one) have "?k \ 1 / k" using lapprox_rat[where x=1 and y=k] by auto have "?k * lb_arctan_horner prec (get_even n) 1 ?ku \ arctan ?k" using arctan_0_1_bounds_round[OF \0 \ real_of_float ?k\ \real_of_float ?k \ 1\] by auto also have "\ \ arctan (1 / k)" using \?k \ 1 / k\ by (rule arctan_monotone') finally have "?k * lb_arctan_horner prec (get_even n) 1 ?ku \ arctan (1 / k)" . } note lb_arctan = this have "pi \ ub_pi n " unfolding ub_pi_def machin_pi Let_def times_float.rep_eq Float_num using lb_arctan[of 239] ub_arctan[of 5] powr_realpow[of 2 2] by (intro mult_left_mono float_plus_up_le float_plus_down_le) (auto intro!: mult_left_mono float_round_down_le float_round_up_le diff_mono) moreover have "lb_pi n \ pi" unfolding lb_pi_def machin_pi Let_def times_float.rep_eq Float_num using lb_arctan[of 5] ub_arctan[of 239] by (intro mult_left_mono float_plus_up_le float_plus_down_le) (auto intro!: mult_left_mono float_round_down_le float_round_up_le diff_mono) ultimately show ?thesis by auto qed lift_definition pi_float_interval::"nat \ float interval" is "\prec. (lb_pi prec, ub_pi prec)" using pi_boundaries by (auto intro: order_trans) lemma lower_pi_float_interval: "lower (pi_float_interval prec) = lb_pi prec" by transfer auto lemma upper_pi_float_interval: "upper (pi_float_interval prec) = ub_pi prec" by transfer auto lemma pi_float_interval: "pi \ set_of (real_interval (pi_float_interval prec))" using pi_boundaries by (auto simp: set_of_eq lower_pi_float_interval upper_pi_float_interval) subsection "Compute arcus tangens in the entire domain" function lb_arctan :: "nat \ float \ float" and ub_arctan :: "nat \ float \ float" where "lb_arctan prec x = (let ub_horner = \ x. float_round_up prec (x * ub_arctan_horner prec (get_odd (prec div 4 + 1)) 1 (float_round_down (Suc prec) (x * x))); lb_horner = \ x. float_round_down prec (x * lb_arctan_horner prec (get_even (prec div 4 + 1)) 1 (float_round_up (Suc prec) (x * x))) in if x < 0 then - ub_arctan prec (-x) else if x \ Float 1 (- 1) then lb_horner x else if x \ Float 1 1 then Float 1 1 * lb_horner (float_divl prec x (float_plus_up prec 1 (ub_sqrt prec (float_plus_up prec 1 (float_round_up prec (x * x)))))) else let inv = float_divr prec 1 x in if inv > 1 then 0 else float_plus_down prec (lb_pi prec * Float 1 (- 1)) ( - ub_horner inv))" | "ub_arctan prec x = (let lb_horner = \ x. float_round_down prec (x * lb_arctan_horner prec (get_even (prec div 4 + 1)) 1 (float_round_up (Suc prec) (x * x))) ; ub_horner = \ x. float_round_up prec (x * ub_arctan_horner prec (get_odd (prec div 4 + 1)) 1 (float_round_down (Suc prec) (x * x))) in if x < 0 then - lb_arctan prec (-x) else if x \ Float 1 (- 1) then ub_horner x else if x \ Float 1 1 then let y = float_divr prec x (float_plus_down (Suc prec) 1 (lb_sqrt prec (float_plus_down prec 1 (float_round_down prec (x * x))))) in if y > 1 then ub_pi prec * Float 1 (- 1) else Float 1 1 * ub_horner y else float_plus_up prec (ub_pi prec * Float 1 (- 1)) ( - lb_horner (float_divl prec 1 x)))" by pat_completeness auto termination by (relation "measure (\ v. let (prec, x) = case_sum id id v in (if x < 0 then 1 else 0))", auto) declare ub_arctan_horner.simps[simp del] declare lb_arctan_horner.simps[simp del] lemma lb_arctan_bound': assumes "0 \ real_of_float x" shows "lb_arctan prec x \ arctan x" proof - have "\ x < 0" and "0 \ x" using \0 \ real_of_float x\ by (auto intro!: truncate_up_le ) let "?ub_horner x" = "x * ub_arctan_horner prec (get_odd (prec div 4 + 1)) 1 (float_round_down (Suc prec) (x * x))" and "?lb_horner x" = "x * lb_arctan_horner prec (get_even (prec div 4 + 1)) 1 (float_round_up (Suc prec) (x * x))" show ?thesis proof (cases "x \ Float 1 (- 1)") case True hence "real_of_float x \ 1" by simp from arctan_0_1_bounds_round[OF \0 \ real_of_float x\ \real_of_float x \ 1\] show ?thesis unfolding lb_arctan.simps Let_def if_not_P[OF \\ x < 0\] if_P[OF True] using \0 \ x\ by (auto intro!: float_round_down_le) next case False hence "0 < real_of_float x" by auto let ?R = "1 + sqrt (1 + real_of_float x * real_of_float x)" let ?sxx = "float_plus_up prec 1 (float_round_up prec (x * x))" let ?fR = "float_plus_up prec 1 (ub_sqrt prec ?sxx)" let ?DIV = "float_divl prec x ?fR" have divisor_gt0: "0 < ?R" by (auto intro: add_pos_nonneg) have "sqrt (1 + x*x) \ sqrt ?sxx" by (auto simp: float_plus_up.rep_eq plus_up_def float_round_up.rep_eq intro!: truncate_up_le) also have "\ \ ub_sqrt prec ?sxx" using bnds_sqrt'[of ?sxx prec] by auto finally have "sqrt (1 + x*x) \ ub_sqrt prec ?sxx" . hence "?R \ ?fR" by (auto simp: float_plus_up.rep_eq plus_up_def intro!: truncate_up_le) hence "0 < ?fR" and "0 < real_of_float ?fR" using \0 < ?R\ by auto have monotone: "?DIV \ x / ?R" proof - have "?DIV \ real_of_float x / ?fR" by (rule float_divl) also have "\ \ x / ?R" by (rule divide_left_mono[OF \?R \ ?fR\ \0 \ real_of_float x\ mult_pos_pos[OF order_less_le_trans[OF divisor_gt0 \?R \ real_of_float ?fR\] divisor_gt0]]) finally show ?thesis . qed show ?thesis proof (cases "x \ Float 1 1") case True have "x \ sqrt (1 + x * x)" using real_sqrt_sum_squares_ge2[where x=1, unfolded numeral_2_eq_2] by auto also note \\ \ (ub_sqrt prec ?sxx)\ finally have "real_of_float x \ ?fR" by (auto simp: float_plus_up.rep_eq plus_up_def intro!: truncate_up_le) moreover have "?DIV \ real_of_float x / ?fR" by (rule float_divl) ultimately have "real_of_float ?DIV \ 1" unfolding divide_le_eq_1_pos[OF \0 < real_of_float ?fR\, symmetric] by auto have "0 \ real_of_float ?DIV" using float_divl_lower_bound[OF \0 \ x\] \0 < ?fR\ unfolding less_eq_float_def by auto from arctan_0_1_bounds_round[OF \0 \ real_of_float (?DIV)\ \real_of_float (?DIV) \ 1\] have "Float 1 1 * ?lb_horner ?DIV \ 2 * arctan ?DIV" by simp also have "\ \ 2 * arctan (x / ?R)" using arctan_monotone'[OF monotone] by (auto intro!: mult_left_mono arctan_monotone') also have "2 * arctan (x / ?R) = arctan x" using arctan_half[symmetric] unfolding numeral_2_eq_2 power_Suc2 power_0 mult_1_left . finally show ?thesis unfolding lb_arctan.simps Let_def if_not_P[OF \\ x < 0\] if_not_P[OF \\ x \ Float 1 (- 1)\] if_P[OF True] by (auto simp: float_round_down.rep_eq intro!: order_trans[OF mult_left_mono[OF truncate_down]]) next case False hence "2 < real_of_float x" by auto hence "1 \ real_of_float x" by auto let "?invx" = "float_divr prec 1 x" have "0 \ arctan x" using arctan_monotone'[OF \0 \ real_of_float x\] using arctan_tan[of 0, unfolded tan_zero] by auto show ?thesis proof (cases "1 < ?invx") case True show ?thesis unfolding lb_arctan.simps Let_def if_not_P[OF \\ x < 0\] if_not_P[OF \\ x \ Float 1 (- 1)\] if_not_P[OF False] if_P[OF True] using \0 \ arctan x\ by auto next case False hence "real_of_float ?invx \ 1" by auto have "0 \ real_of_float ?invx" by (rule order_trans[OF _ float_divr]) (auto simp add: \0 \ real_of_float x\) have "1 / x \ 0" and "0 < 1 / x" using \0 < real_of_float x\ by auto have "arctan (1 / x) \ arctan ?invx" unfolding one_float.rep_eq[symmetric] by (rule arctan_monotone', rule float_divr) also have "\ \ ?ub_horner ?invx" using arctan_0_1_bounds_round[OF \0 \ real_of_float ?invx\ \real_of_float ?invx \ 1\] by (auto intro!: float_round_up_le) also note float_round_up finally have "pi / 2 - float_round_up prec (?ub_horner ?invx) \ arctan x" using \0 \ arctan x\ arctan_inverse[OF \1 / x \ 0\] unfolding sgn_pos[OF \0 < 1 / real_of_float x\] le_diff_eq by auto moreover have "lb_pi prec * Float 1 (- 1) \ pi / 2" unfolding Float_num times_divide_eq_right mult_1_left using pi_boundaries by simp ultimately show ?thesis unfolding lb_arctan.simps Let_def if_not_P[OF \\ x < 0\] if_not_P[OF \\ x \ Float 1 (- 1)\] if_not_P[OF \\ x \ Float 1 1\] if_not_P[OF False] by (auto intro!: float_plus_down_le) qed qed qed qed lemma ub_arctan_bound': assumes "0 \ real_of_float x" shows "arctan x \ ub_arctan prec x" proof - have "\ x < 0" and "0 \ x" using \0 \ real_of_float x\ by auto let "?ub_horner x" = "float_round_up prec (x * ub_arctan_horner prec (get_odd (prec div 4 + 1)) 1 (float_round_down (Suc prec) (x * x)))" let "?lb_horner x" = "float_round_down prec (x * lb_arctan_horner prec (get_even (prec div 4 + 1)) 1 (float_round_up (Suc prec) (x * x)))" show ?thesis proof (cases "x \ Float 1 (- 1)") case True hence "real_of_float x \ 1" by auto show ?thesis unfolding ub_arctan.simps Let_def if_not_P[OF \\ x < 0\] if_P[OF True] using arctan_0_1_bounds_round[OF \0 \ real_of_float x\ \real_of_float x \ 1\] by (auto intro!: float_round_up_le) next case False hence "0 < real_of_float x" by auto let ?R = "1 + sqrt (1 + real_of_float x * real_of_float x)" let ?sxx = "float_plus_down prec 1 (float_round_down prec (x * x))" let ?fR = "float_plus_down (Suc prec) 1 (lb_sqrt prec ?sxx)" let ?DIV = "float_divr prec x ?fR" have sqr_ge0: "0 \ 1 + real_of_float x * real_of_float x" using sum_power2_ge_zero[of 1 "real_of_float x", unfolded numeral_2_eq_2] by auto hence "0 \ real_of_float (1 + x*x)" by auto hence divisor_gt0: "0 < ?R" by (auto intro: add_pos_nonneg) have "lb_sqrt prec ?sxx \ sqrt ?sxx" using bnds_sqrt'[of ?sxx] by auto also have "\ \ sqrt (1 + x*x)" by (auto simp: float_plus_down.rep_eq plus_down_def float_round_down.rep_eq truncate_down_le) finally have "lb_sqrt prec ?sxx \ sqrt (1 + x*x)" . hence "?fR \ ?R" by (auto simp: float_plus_down.rep_eq plus_down_def truncate_down_le) have "0 < real_of_float ?fR" by (auto simp: float_plus_down.rep_eq plus_down_def float_round_down.rep_eq intro!: truncate_down_ge1 lb_sqrt_lower_bound order_less_le_trans[OF zero_less_one] truncate_down_nonneg add_nonneg_nonneg) have monotone: "x / ?R \ (float_divr prec x ?fR)" proof - from divide_left_mono[OF \?fR \ ?R\ \0 \ real_of_float x\ mult_pos_pos[OF divisor_gt0 \0 < real_of_float ?fR\]] have "x / ?R \ x / ?fR" . also have "\ \ ?DIV" by (rule float_divr) finally show ?thesis . qed show ?thesis proof (cases "x \ Float 1 1") case True show ?thesis proof (cases "?DIV > 1") case True have "pi / 2 \ ub_pi prec * Float 1 (- 1)" unfolding Float_num times_divide_eq_right mult_1_left using pi_boundaries by auto from order_less_le_trans[OF arctan_ubound this, THEN less_imp_le] show ?thesis unfolding ub_arctan.simps Let_def if_not_P[OF \\ x < 0\] if_not_P[OF \\ x \ Float 1 (- 1)\] if_P[OF \x \ Float 1 1\] if_P[OF True] . next case False hence "real_of_float ?DIV \ 1" by auto have "0 \ x / ?R" using \0 \ real_of_float x\ \0 < ?R\ unfolding zero_le_divide_iff by auto hence "0 \ real_of_float ?DIV" using monotone by (rule order_trans) have "arctan x = 2 * arctan (x / ?R)" using arctan_half unfolding numeral_2_eq_2 power_Suc2 power_0 mult_1_left . also have "\ \ 2 * arctan (?DIV)" using arctan_monotone'[OF monotone] by (auto intro!: mult_left_mono) also have "\ \ (Float 1 1 * ?ub_horner ?DIV)" unfolding Float_num using arctan_0_1_bounds_round[OF \0 \ real_of_float ?DIV\ \real_of_float ?DIV \ 1\] by (auto intro!: float_round_up_le) finally show ?thesis unfolding ub_arctan.simps Let_def if_not_P[OF \\ x < 0\] if_not_P[OF \\ x \ Float 1 (- 1)\] if_P[OF \x \ Float 1 1\] if_not_P[OF False] . qed next case False hence "2 < real_of_float x" by auto hence "1 \ real_of_float x" by auto hence "0 < real_of_float x" by auto hence "0 < x" by auto let "?invx" = "float_divl prec 1 x" have "0 \ arctan x" using arctan_monotone'[OF \0 \ real_of_float x\] and arctan_tan[of 0, unfolded tan_zero] by auto have "real_of_float ?invx \ 1" unfolding less_float_def by (rule order_trans[OF float_divl]) (auto simp add: \1 \ real_of_float x\ divide_le_eq_1_pos[OF \0 < real_of_float x\]) have "0 \ real_of_float ?invx" using \0 < x\ by (intro float_divl_lower_bound) auto have "1 / x \ 0" and "0 < 1 / x" using \0 < real_of_float x\ by auto have "(?lb_horner ?invx) \ arctan (?invx)" using arctan_0_1_bounds_round[OF \0 \ real_of_float ?invx\ \real_of_float ?invx \ 1\] by (auto intro!: float_round_down_le) also have "\ \ arctan (1 / x)" unfolding one_float.rep_eq[symmetric] by (rule arctan_monotone') (rule float_divl) finally have "arctan x \ pi / 2 - (?lb_horner ?invx)" using \0 \ arctan x\ arctan_inverse[OF \1 / x \ 0\] unfolding sgn_pos[OF \0 < 1 / x\] le_diff_eq by auto moreover have "pi / 2 \ ub_pi prec * Float 1 (- 1)" unfolding Float_num times_divide_eq_right mult_1_right using pi_boundaries by auto ultimately show ?thesis unfolding ub_arctan.simps Let_def if_not_P[OF \\ x < 0\] if_not_P[OF \\ x \ Float 1 (- 1)\] if_not_P[OF False] by (auto intro!: float_round_up_le float_plus_up_le) qed qed qed lemma arctan_boundaries: "arctan x \ {(lb_arctan prec x) .. (ub_arctan prec x)}" proof (cases "0 \ x") case True hence "0 \ real_of_float x" by auto show ?thesis using ub_arctan_bound'[OF \0 \ real_of_float x\] lb_arctan_bound'[OF \0 \ real_of_float x\] unfolding atLeastAtMost_iff by auto next case False let ?mx = "-x" from False have "x < 0" and "0 \ real_of_float ?mx" by auto hence bounds: "lb_arctan prec ?mx \ arctan ?mx \ arctan ?mx \ ub_arctan prec ?mx" using ub_arctan_bound'[OF \0 \ real_of_float ?mx\] lb_arctan_bound'[OF \0 \ real_of_float ?mx\] by auto show ?thesis unfolding minus_float.rep_eq arctan_minus lb_arctan.simps[where x=x] ub_arctan.simps[where x=x] Let_def if_P[OF \x < 0\] unfolding atLeastAtMost_iff using bounds[unfolded minus_float.rep_eq arctan_minus] by (simp add: arctan_minus) qed lemma bnds_arctan: "\ (x::real) lx ux. (l, u) = (lb_arctan prec lx, ub_arctan prec ux) \ x \ {lx .. ux} \ l \ arctan x \ arctan x \ u" proof (rule allI, rule allI, rule allI, rule impI) fix x :: real fix lx ux assume "(l, u) = (lb_arctan prec lx, ub_arctan prec ux) \ x \ {lx .. ux}" hence l: "lb_arctan prec lx = l " and u: "ub_arctan prec ux = u" and x: "x \ {lx .. ux}" by auto show "l \ arctan x \ arctan x \ u" proof show "l \ arctan x" proof - from arctan_boundaries[of lx prec, unfolded l] have "l \ arctan lx" by (auto simp del: lb_arctan.simps) also have "\ \ arctan x" using x by (auto intro: arctan_monotone') finally show ?thesis . qed show "arctan x \ u" proof - have "arctan x \ arctan ux" using x by (auto intro: arctan_monotone') also have "\ \ u" using arctan_boundaries[of ux prec, unfolded u] by (auto simp del: ub_arctan.simps) finally show ?thesis . qed qed qed lemmas [simp del] = lb_arctan.simps ub_arctan.simps lemma lb_arctan: "arctan (real_of_float x) \ y \ real_of_float (lb_arctan prec x) \ y" and ub_arctan: "y \ arctan x \ y \ ub_arctan prec x" for x::float and y::real using arctan_boundaries[of x prec] by auto lift_definition arctan_float_interval :: "nat \ float interval \ float interval" is "\prec. \(lx, ux). (lb_arctan prec lx, ub_arctan prec ux)" by (auto intro!: lb_arctan ub_arctan arctan_monotone') lemma lower_arctan_float_interval: "lower (arctan_float_interval p x) = lb_arctan p (lower x)" by transfer auto lemma upper_arctan_float_interval: "upper (arctan_float_interval p x) = ub_arctan p (upper x)" by transfer auto lemma arctan_float_interval: "arctan ` set_of (real_interval x) \ set_of (real_interval (arctan_float_interval p x))" by (auto simp: set_of_eq lower_arctan_float_interval upper_arctan_float_interval intro!: lb_arctan ub_arctan arctan_monotone') lemma arctan_float_intervalI: "arctan x \\<^sub>r arctan_float_interval p X" if "x \\<^sub>r X" using arctan_float_interval[of X p] that by auto section "Sinus and Cosinus" subsection "Compute the cosinus and sinus series" fun ub_sin_cos_aux :: "nat \ nat \ nat \ nat \ float \ float" and lb_sin_cos_aux :: "nat \ nat \ nat \ nat \ float \ float" where "ub_sin_cos_aux prec 0 i k x = 0" | "ub_sin_cos_aux prec (Suc n) i k x = float_plus_up prec (rapprox_rat prec 1 k) (- float_round_down prec (x * (lb_sin_cos_aux prec n (i + 2) (k * i * (i + 1)) x)))" | "lb_sin_cos_aux prec 0 i k x = 0" | "lb_sin_cos_aux prec (Suc n) i k x = float_plus_down prec (lapprox_rat prec 1 k) (- float_round_up prec (x * (ub_sin_cos_aux prec n (i + 2) (k * i * (i + 1)) x)))" lemma cos_aux: shows "(lb_sin_cos_aux prec n 1 1 (x * x)) \ (\ i=0.. i=0.. (ub_sin_cos_aux prec n 1 1 (x * x))" (is "?ub") proof - have "0 \ real_of_float (x * x)" by auto let "?f n" = "fact (2 * n) :: nat" have f_eq: "?f (Suc n) = ?f n * ((\i. i + 2) ^^ n) 1 * (((\i. i + 2) ^^ n) 1 + 1)" for n proof - have "\m. ((\i. i + 2) ^^ n) m = m + 2 * n" by (induct n) auto then show ?thesis by auto qed from horner_bounds[where lb="lb_sin_cos_aux prec" and ub="ub_sin_cos_aux prec" and j'=0, OF \0 \ real_of_float (x * x)\ f_eq lb_sin_cos_aux.simps ub_sin_cos_aux.simps] show ?lb and ?ub by (auto simp add: power_mult power2_eq_square[of "real_of_float x"]) qed lemma lb_sin_cos_aux_zero_le_one: "lb_sin_cos_aux prec n i j 0 \ 1" by (cases j n rule: nat.exhaust[case_product nat.exhaust]) (auto intro!: float_plus_down_le order_trans[OF lapprox_rat]) lemma one_le_ub_sin_cos_aux: "odd n \ 1 \ ub_sin_cos_aux prec n i (Suc 0) 0" by (cases n) (auto intro!: float_plus_up_le order_trans[OF _ rapprox_rat]) lemma cos_boundaries: assumes "0 \ real_of_float x" and "x \ pi / 2" shows "cos x \ {(lb_sin_cos_aux prec (get_even n) 1 1 (x * x)) .. (ub_sin_cos_aux prec (get_odd n) 1 1 (x * x))}" proof (cases "real_of_float x = 0") case False hence "real_of_float x \ 0" by auto hence "0 < x" and "0 < real_of_float x" using \0 \ real_of_float x\ by auto have "0 < x * x" using \0 < x\ by simp have morph_to_if_power: "(\ i=0.. i = 0 ..< 2 * n. (if even(i) then ((- 1) ^ (i div 2))/((fact i)) else 0) * x ^ i)" (is "?sum = ?ifsum") for x n proof - have "?sum = ?sum + (\ j = 0 ..< n. 0)" by auto also have "\ = (\ j = 0 ..< n. (- 1) ^ ((2 * j) div 2) / ((fact (2 * j))) * x ^(2 * j)) + (\ j = 0 ..< n. 0)" by auto also have "\ = (\ i = 0 ..< 2 * n. if even i then (- 1) ^ (i div 2) / ((fact i)) * x ^ i else 0)" unfolding sum_split_even_odd atLeast0LessThan .. also have "\ = (\ i = 0 ..< 2 * n. (if even i then (- 1) ^ (i div 2) / ((fact i)) else 0) * x ^ i)" by (rule sum.cong) auto finally show ?thesis . qed { fix n :: nat assume "0 < n" hence "0 < 2 * n" by auto obtain t where "0 < t" and "t < real_of_float x" and cos_eq: "cos x = (\ i = 0 ..< 2 * n. (if even(i) then ((- 1) ^ (i div 2))/((fact i)) else 0) * (real_of_float x) ^ i) + (cos (t + 1/2 * (2 * n) * pi) / (fact (2*n))) * (real_of_float x)^(2*n)" (is "_ = ?SUM + ?rest / ?fact * ?pow") using Maclaurin_cos_expansion2[OF \0 < real_of_float x\ \0 < 2 * n\] unfolding cos_coeff_def atLeast0LessThan by auto have "cos t * (- 1) ^ n = cos t * cos (n * pi) + sin t * sin (n * pi)" by auto also have "\ = cos (t + n * pi)" by (simp add: cos_add) also have "\ = ?rest" by auto finally have "cos t * (- 1) ^ n = ?rest" . moreover have "t \ pi / 2" using \t < real_of_float x\ and \x \ pi / 2\ by auto hence "0 \ cos t" using \0 < t\ and cos_ge_zero by auto ultimately have even: "even n \ 0 \ ?rest" and odd: "odd n \ 0 \ - ?rest " by auto have "0 < ?fact" by auto have "0 < ?pow" using \0 < real_of_float x\ by auto { assume "even n" have "(lb_sin_cos_aux prec n 1 1 (x * x)) \ ?SUM" unfolding morph_to_if_power[symmetric] using cos_aux by auto also have "\ \ cos x" proof - from even[OF \even n\] \0 < ?fact\ \0 < ?pow\ have "0 \ (?rest / ?fact) * ?pow" by simp thus ?thesis unfolding cos_eq by auto qed finally have "(lb_sin_cos_aux prec n 1 1 (x * x)) \ cos x" . } note lb = this { assume "odd n" have "cos x \ ?SUM" proof - from \0 < ?fact\ and \0 < ?pow\ and odd[OF \odd n\] have "0 \ (- ?rest) / ?fact * ?pow" by (metis mult_nonneg_nonneg divide_nonneg_pos less_imp_le) thus ?thesis unfolding cos_eq by auto qed also have "\ \ (ub_sin_cos_aux prec n 1 1 (x * x))" unfolding morph_to_if_power[symmetric] using cos_aux by auto finally have "cos x \ (ub_sin_cos_aux prec n 1 1 (x * x))" . } note ub = this and lb } note ub = this(1) and lb = this(2) have "cos x \ (ub_sin_cos_aux prec (get_odd n) 1 1 (x * x))" using ub[OF odd_pos[OF get_odd] get_odd] . moreover have "(lb_sin_cos_aux prec (get_even n) 1 1 (x * x)) \ cos x" proof (cases "0 < get_even n") case True show ?thesis using lb[OF True get_even] . next case False hence "get_even n = 0" by auto have "- (pi / 2) \ x" by (rule order_trans[OF _ \0 < real_of_float x\[THEN less_imp_le]]) auto with \x \ pi / 2\ show ?thesis unfolding \get_even n = 0\ lb_sin_cos_aux.simps minus_float.rep_eq zero_float.rep_eq using cos_ge_zero by auto qed ultimately show ?thesis by auto next case True hence "x = 0" by (simp add: real_of_float_eq) thus ?thesis using lb_sin_cos_aux_zero_le_one one_le_ub_sin_cos_aux by simp qed lemma sin_aux: assumes "0 \ real_of_float x" shows "(x * lb_sin_cos_aux prec n 2 1 (x * x)) \ (\ i=0.. i=0.. (x * ub_sin_cos_aux prec n 2 1 (x * x))" (is "?ub") proof - have "0 \ real_of_float (x * x)" by auto let "?f n" = "fact (2 * n + 1) :: nat" have f_eq: "?f (Suc n) = ?f n * ((\i. i + 2) ^^ n) 2 * (((\i. i + 2) ^^ n) 2 + 1)" for n proof - have F: "\m. ((\i. i + 2) ^^ n) m = m + 2 * n" by (induct n) auto show ?thesis unfolding F by auto qed from horner_bounds[where lb="lb_sin_cos_aux prec" and ub="ub_sin_cos_aux prec" and j'=0, OF \0 \ real_of_float (x * x)\ f_eq lb_sin_cos_aux.simps ub_sin_cos_aux.simps] show "?lb" and "?ub" using \0 \ real_of_float x\ apply (simp_all only: power_add power_one_right mult.assoc[symmetric] sum_distrib_right[symmetric]) apply (simp_all only: mult.commute[where 'a=real] of_nat_fact) apply (auto intro!: mult_left_mono simp add: power_mult power2_eq_square[of "real_of_float x"]) done qed lemma sin_boundaries: assumes "0 \ real_of_float x" and "x \ pi / 2" shows "sin x \ {(x * lb_sin_cos_aux prec (get_even n) 2 1 (x * x)) .. (x * ub_sin_cos_aux prec (get_odd n) 2 1 (x * x))}" proof (cases "real_of_float x = 0") case False hence "real_of_float x \ 0" by auto hence "0 < x" and "0 < real_of_float x" using \0 \ real_of_float x\ by auto have "0 < x * x" using \0 < x\ by simp have sum_morph: "(\j = 0 ..< n. (- 1) ^ (((2 * j + 1) - Suc 0) div 2) / ((fact (2 * j + 1))) * x ^(2 * j + 1)) = (\ i = 0 ..< 2 * n. (if even(i) then 0 else ((- 1) ^ ((i - Suc 0) div 2))/((fact i))) * x ^ i)" (is "?SUM = _") for x :: real and n proof - have pow: "!!i. x ^ (2 * i + 1) = x * x ^ (2 * i)" by auto have "?SUM = (\ j = 0 ..< n. 0) + ?SUM" by auto also have "\ = (\ i = 0 ..< 2 * n. if even i then 0 else (- 1) ^ ((i - Suc 0) div 2) / ((fact i)) * x ^ i)" unfolding sum_split_even_odd atLeast0LessThan .. also have "\ = (\ i = 0 ..< 2 * n. (if even i then 0 else (- 1) ^ ((i - Suc 0) div 2) / ((fact i))) * x ^ i)" by (rule sum.cong) auto finally show ?thesis . qed { fix n :: nat assume "0 < n" hence "0 < 2 * n + 1" by auto obtain t where "0 < t" and "t < real_of_float x" and sin_eq: "sin x = (\ i = 0 ..< 2 * n + 1. (if even(i) then 0 else ((- 1) ^ ((i - Suc 0) div 2))/((fact i))) * (real_of_float x) ^ i) + (sin (t + 1/2 * (2 * n + 1) * pi) / (fact (2*n + 1))) * (real_of_float x)^(2*n + 1)" (is "_ = ?SUM + ?rest / ?fact * ?pow") using Maclaurin_sin_expansion3[OF \0 < 2 * n + 1\ \0 < real_of_float x\] unfolding sin_coeff_def atLeast0LessThan by auto have "?rest = cos t * (- 1) ^ n" unfolding sin_add cos_add of_nat_add distrib_right distrib_left by auto moreover have "t \ pi / 2" using \t < real_of_float x\ and \x \ pi / 2\ by auto hence "0 \ cos t" using \0 < t\ and cos_ge_zero by auto ultimately have even: "even n \ 0 \ ?rest" and odd: "odd n \ 0 \ - ?rest" by auto have "0 < ?fact" by (simp del: fact_Suc) have "0 < ?pow" using \0 < real_of_float x\ by (rule zero_less_power) { assume "even n" have "(x * lb_sin_cos_aux prec n 2 1 (x * x)) \ (\ i = 0 ..< 2 * n. (if even(i) then 0 else ((- 1) ^ ((i - Suc 0) div 2))/((fact i))) * (real_of_float x) ^ i)" using sin_aux[OF \0 \ real_of_float x\] unfolding sum_morph[symmetric] by auto also have "\ \ ?SUM" by auto also have "\ \ sin x" proof - from even[OF \even n\] \0 < ?fact\ \0 < ?pow\ have "0 \ (?rest / ?fact) * ?pow" by simp thus ?thesis unfolding sin_eq by auto qed finally have "(x * lb_sin_cos_aux prec n 2 1 (x * x)) \ sin x" . } note lb = this { assume "odd n" have "sin x \ ?SUM" proof - from \0 < ?fact\ and \0 < ?pow\ and odd[OF \odd n\] have "0 \ (- ?rest) / ?fact * ?pow" by (metis mult_nonneg_nonneg divide_nonneg_pos less_imp_le) thus ?thesis unfolding sin_eq by auto qed also have "\ \ (\ i = 0 ..< 2 * n. (if even(i) then 0 else ((- 1) ^ ((i - Suc 0) div 2))/((fact i))) * (real_of_float x) ^ i)" by auto also have "\ \ (x * ub_sin_cos_aux prec n 2 1 (x * x))" using sin_aux[OF \0 \ real_of_float x\] unfolding sum_morph[symmetric] by auto finally have "sin x \ (x * ub_sin_cos_aux prec n 2 1 (x * x))" . } note ub = this and lb } note ub = this(1) and lb = this(2) have "sin x \ (x * ub_sin_cos_aux prec (get_odd n) 2 1 (x * x))" using ub[OF odd_pos[OF get_odd] get_odd] . moreover have "(x * lb_sin_cos_aux prec (get_even n) 2 1 (x * x)) \ sin x" proof (cases "0 < get_even n") case True show ?thesis using lb[OF True get_even] . next case False hence "get_even n = 0" by auto with \x \ pi / 2\ \0 \ real_of_float x\ show ?thesis unfolding \get_even n = 0\ ub_sin_cos_aux.simps minus_float.rep_eq using sin_ge_zero by auto qed ultimately show ?thesis by auto next case True show ?thesis proof (cases "n = 0") case True thus ?thesis unfolding \n = 0\ get_even_def get_odd_def using \real_of_float x = 0\ lapprox_rat[where x="-1" and y=1] by auto next case False with not0_implies_Suc obtain m where "n = Suc m" by blast thus ?thesis unfolding \n = Suc m\ get_even_def get_odd_def using \real_of_float x = 0\ rapprox_rat[where x=1 and y=1] lapprox_rat[where x=1 and y=1] by (cases "even (Suc m)") auto qed qed subsection "Compute the cosinus in the entire domain" definition lb_cos :: "nat \ float \ float" where "lb_cos prec x = (let horner = \ x. lb_sin_cos_aux prec (get_even (prec div 4 + 1)) 1 1 (x * x) ; half = \ x. if x < 0 then - 1 else float_plus_down prec (Float 1 1 * x * x) (- 1) in if x < Float 1 (- 1) then horner x else if x < 1 then half (horner (x * Float 1 (- 1))) else half (half (horner (x * Float 1 (- 2)))))" definition ub_cos :: "nat \ float \ float" where "ub_cos prec x = (let horner = \ x. ub_sin_cos_aux prec (get_odd (prec div 4 + 1)) 1 1 (x * x) ; half = \ x. float_plus_up prec (Float 1 1 * x * x) (- 1) in if x < Float 1 (- 1) then horner x else if x < 1 then half (horner (x * Float 1 (- 1))) else half (half (horner (x * Float 1 (- 2)))))" lemma lb_cos: assumes "0 \ real_of_float x" and "x \ pi" shows "cos x \ {(lb_cos prec x) .. (ub_cos prec x)}" (is "?cos x \ {(?lb x) .. (?ub x) }") proof - have x_half[symmetric]: "cos x = 2 * cos (x / 2) * cos (x / 2) - 1" for x :: real proof - have "cos x = cos (x / 2 + x / 2)" by auto also have "\ = cos (x / 2) * cos (x / 2) + sin (x / 2) * sin (x / 2) - sin (x / 2) * sin (x / 2) + cos (x / 2) * cos (x / 2) - 1" unfolding cos_add by auto also have "\ = 2 * cos (x / 2) * cos (x / 2) - 1" by algebra finally show ?thesis . qed have "\ x < 0" using \0 \ real_of_float x\ by auto let "?ub_horner x" = "ub_sin_cos_aux prec (get_odd (prec div 4 + 1)) 1 1 (x * x)" let "?lb_horner x" = "lb_sin_cos_aux prec (get_even (prec div 4 + 1)) 1 1 (x * x)" let "?ub_half x" = "float_plus_up prec (Float 1 1 * x * x) (- 1)" let "?lb_half x" = "if x < 0 then - 1 else float_plus_down prec (Float 1 1 * x * x) (- 1)" show ?thesis proof (cases "x < Float 1 (- 1)") case True hence "x \ pi / 2" using pi_ge_two by auto show ?thesis unfolding lb_cos_def[where x=x] ub_cos_def[where x=x] if_not_P[OF \\ x < 0\] if_P[OF \x < Float 1 (- 1)\] Let_def using cos_boundaries[OF \0 \ real_of_float x\ \x \ pi / 2\] . next case False { fix y x :: float let ?x2 = "(x * Float 1 (- 1))" assume "y \ cos ?x2" and "-pi \ x" and "x \ pi" hence "- (pi / 2) \ ?x2" and "?x2 \ pi / 2" using pi_ge_two unfolding Float_num by auto hence "0 \ cos ?x2" by (rule cos_ge_zero) have "(?lb_half y) \ cos x" proof (cases "y < 0") case True show ?thesis using cos_ge_minus_one unfolding if_P[OF True] by auto next case False hence "0 \ real_of_float y" by auto from mult_mono[OF \y \ cos ?x2\ \y \ cos ?x2\ \0 \ cos ?x2\ this] have "real_of_float y * real_of_float y \ cos ?x2 * cos ?x2" . hence "2 * real_of_float y * real_of_float y \ 2 * cos ?x2 * cos ?x2" by auto hence "2 * real_of_float y * real_of_float y - 1 \ 2 * cos (x / 2) * cos (x / 2) - 1" unfolding Float_num by auto thus ?thesis unfolding if_not_P[OF False] x_half Float_num by (auto intro!: float_plus_down_le) qed } note lb_half = this { fix y x :: float let ?x2 = "(x * Float 1 (- 1))" assume ub: "cos ?x2 \ y" and "- pi \ x" and "x \ pi" hence "- (pi / 2) \ ?x2" and "?x2 \ pi / 2" using pi_ge_two unfolding Float_num by auto hence "0 \ cos ?x2" by (rule cos_ge_zero) have "cos x \ (?ub_half y)" proof - have "0 \ real_of_float y" using \0 \ cos ?x2\ ub by (rule order_trans) from mult_mono[OF ub ub this \0 \ cos ?x2\] have "cos ?x2 * cos ?x2 \ real_of_float y * real_of_float y" . hence "2 * cos ?x2 * cos ?x2 \ 2 * real_of_float y * real_of_float y" by auto hence "2 * cos (x / 2) * cos (x / 2) - 1 \ 2 * real_of_float y * real_of_float y - 1" unfolding Float_num by auto thus ?thesis unfolding x_half Float_num by (auto intro!: float_plus_up_le) qed } note ub_half = this let ?x2 = "x * Float 1 (- 1)" let ?x4 = "x * Float 1 (- 1) * Float 1 (- 1)" have "-pi \ x" using pi_ge_zero[THEN le_imp_neg_le, unfolded minus_zero] \0 \ real_of_float x\ by (rule order_trans) show ?thesis proof (cases "x < 1") case True hence "real_of_float x \ 1" by auto have "0 \ real_of_float ?x2" and "?x2 \ pi / 2" using pi_ge_two \0 \ real_of_float x\ using assms by auto from cos_boundaries[OF this] have lb: "(?lb_horner ?x2) \ ?cos ?x2" and ub: "?cos ?x2 \ (?ub_horner ?x2)" by auto have "(?lb x) \ ?cos x" proof - from lb_half[OF lb \-pi \ x\ \x \ pi\] show ?thesis unfolding lb_cos_def[where x=x] Let_def using \\ x < 0\ \\ x < Float 1 (- 1)\ \x < 1\ by auto qed moreover have "?cos x \ (?ub x)" proof - from ub_half[OF ub \-pi \ x\ \x \ pi\] show ?thesis unfolding ub_cos_def[where x=x] Let_def using \\ x < 0\ \\ x < Float 1 (- 1)\ \x < 1\ by auto qed ultimately show ?thesis by auto next case False have "0 \ real_of_float ?x4" and "?x4 \ pi / 2" using pi_ge_two \0 \ real_of_float x\ \x \ pi\ unfolding Float_num by auto from cos_boundaries[OF this] have lb: "(?lb_horner ?x4) \ ?cos ?x4" and ub: "?cos ?x4 \ (?ub_horner ?x4)" by auto have eq_4: "?x2 * Float 1 (- 1) = x * Float 1 (- 2)" by (auto simp: real_of_float_eq) have "(?lb x) \ ?cos x" proof - have "-pi \ ?x2" and "?x2 \ pi" using pi_ge_two \0 \ real_of_float x\ \x \ pi\ by auto from lb_half[OF lb_half[OF lb this] \-pi \ x\ \x \ pi\, unfolded eq_4] show ?thesis unfolding lb_cos_def[where x=x] if_not_P[OF \\ x < 0\] if_not_P[OF \\ x < Float 1 (- 1)\] if_not_P[OF \\ x < 1\] Let_def . qed moreover have "?cos x \ (?ub x)" proof - have "-pi \ ?x2" and "?x2 \ pi" using pi_ge_two \0 \ real_of_float x\ \ x \ pi\ by auto from ub_half[OF ub_half[OF ub this] \-pi \ x\ \x \ pi\, unfolded eq_4] show ?thesis unfolding ub_cos_def[where x=x] if_not_P[OF \\ x < 0\] if_not_P[OF \\ x < Float 1 (- 1)\] if_not_P[OF \\ x < 1\] Let_def . qed ultimately show ?thesis by auto qed qed qed lemma lb_cos_minus: assumes "-pi \ x" and "real_of_float x \ 0" shows "cos (real_of_float(-x)) \ {(lb_cos prec (-x)) .. (ub_cos prec (-x))}" proof - have "0 \ real_of_float (-x)" and "(-x) \ pi" using \-pi \ x\ \real_of_float x \ 0\ by auto from lb_cos[OF this] show ?thesis . qed definition bnds_cos :: "nat \ float \ float \ float * float" where "bnds_cos prec lx ux = (let lpi = float_round_down prec (lb_pi prec) ; upi = float_round_up prec (ub_pi prec) ; k = floor_fl (float_divr prec (lx + lpi) (2 * lpi)) ; lx = float_plus_down prec lx (- k * 2 * (if k < 0 then lpi else upi)) ; ux = float_plus_up prec ux (- k * 2 * (if k < 0 then upi else lpi)) in if - lpi \ lx \ ux \ 0 then (lb_cos prec (-lx), ub_cos prec (-ux)) else if 0 \ lx \ ux \ lpi then (lb_cos prec ux, ub_cos prec lx) else if - lpi \ lx \ ux \ lpi then (min (lb_cos prec (-lx)) (lb_cos prec ux), Float 1 0) else if 0 \ lx \ ux \ 2 * lpi then (Float (- 1) 0, max (ub_cos prec lx) (ub_cos prec (- (ux - 2 * lpi)))) else if -2 * lpi \ lx \ ux \ 0 then (Float (- 1) 0, max (ub_cos prec (lx + 2 * lpi)) (ub_cos prec (-ux))) else (Float (- 1) 0, Float 1 0))" lemma floor_int: obtains k :: int where "real_of_int k = (floor_fl f)" by (simp add: floor_fl_def) lemma cos_periodic_nat[simp]: fixes n :: nat shows "cos (x + n * (2 * pi)) = cos x" proof (induct n arbitrary: x) case 0 then show ?case by simp next case (Suc n) have split_pi_off: "x + (Suc n) * (2 * pi) = (x + n * (2 * pi)) + 2 * pi" unfolding Suc_eq_plus1 of_nat_add of_int_1 distrib_right by auto show ?case unfolding split_pi_off using Suc by auto qed lemma cos_periodic_int[simp]: fixes i :: int shows "cos (x + i * (2 * pi)) = cos x" proof (cases "0 \ i") case True hence i_nat: "real_of_int i = nat i" by auto show ?thesis unfolding i_nat by auto next case False hence i_nat: "i = - real (nat (-i))" by auto have "cos x = cos (x + i * (2 * pi) - i * (2 * pi))" by auto also have "\ = cos (x + i * (2 * pi))" unfolding i_nat mult_minus_left diff_minus_eq_add by (rule cos_periodic_nat) finally show ?thesis by auto qed lemma bnds_cos: "\(x::real) lx ux. (l, u) = bnds_cos prec lx ux \ x \ {lx .. ux} \ l \ cos x \ cos x \ u" proof (rule allI | rule impI | erule conjE)+ fix x :: real fix lx ux assume bnds: "(l, u) = bnds_cos prec lx ux" and x: "x \ {lx .. ux}" let ?lpi = "float_round_down prec (lb_pi prec)" let ?upi = "float_round_up prec (ub_pi prec)" let ?k = "floor_fl (float_divr prec (lx + ?lpi) (2 * ?lpi))" let ?lx2 = "(- ?k * 2 * (if ?k < 0 then ?lpi else ?upi))" let ?ux2 = "(- ?k * 2 * (if ?k < 0 then ?upi else ?lpi))" let ?lx = "float_plus_down prec lx ?lx2" let ?ux = "float_plus_up prec ux ?ux2" obtain k :: int where k: "k = real_of_float ?k" by (rule floor_int) have upi: "pi \ ?upi" and lpi: "?lpi \ pi" using float_round_up[of "ub_pi prec" prec] pi_boundaries[of prec] float_round_down[of prec "lb_pi prec"] by auto hence "lx + ?lx2 \ x - k * (2 * pi) \ x - k * (2 * pi) \ ux + ?ux2" using x by (cases "k = 0") (auto intro!: add_mono simp add: k [symmetric] uminus_add_conv_diff [symmetric] simp del: uminus_add_conv_diff) hence "?lx \ x - k * (2 * pi) \ x - k * (2 * pi) \ ?ux" by (auto intro!: float_plus_down_le float_plus_up_le) note lx = this[THEN conjunct1] and ux = this[THEN conjunct2] hence lx_less_ux: "?lx \ real_of_float ?ux" by (rule order_trans) { assume "- ?lpi \ ?lx" and x_le_0: "x - k * (2 * pi) \ 0" with lpi[THEN le_imp_neg_le] lx have pi_lx: "- pi \ ?lx" and lx_0: "real_of_float ?lx \ 0" by simp_all have "(lb_cos prec (- ?lx)) \ cos (real_of_float (- ?lx))" using lb_cos_minus[OF pi_lx lx_0] by simp also have "\ \ cos (x + (-k) * (2 * pi))" using cos_monotone_minus_pi_0'[OF pi_lx lx x_le_0] by (simp only: uminus_float.rep_eq of_int_minus cos_minus mult_minus_left) simp finally have "(lb_cos prec (- ?lx)) \ cos x" unfolding cos_periodic_int . } note negative_lx = this { assume "0 \ ?lx" and pi_x: "x - k * (2 * pi) \ pi" with lx have pi_lx: "?lx \ pi" and lx_0: "0 \ real_of_float ?lx" by auto have "cos (x + (-k) * (2 * pi)) \ cos ?lx" using cos_monotone_0_pi_le[OF lx_0 lx pi_x] by (simp only: of_int_minus cos_minus mult_minus_left) simp also have "\ \ (ub_cos prec ?lx)" using lb_cos[OF lx_0 pi_lx] by simp finally have "cos x \ (ub_cos prec ?lx)" unfolding cos_periodic_int . } note positive_lx = this { assume pi_x: "- pi \ x - k * (2 * pi)" and "?ux \ 0" with ux have pi_ux: "- pi \ ?ux" and ux_0: "real_of_float ?ux \ 0" by simp_all have "cos (x + (-k) * (2 * pi)) \ cos (real_of_float (- ?ux))" using cos_monotone_minus_pi_0'[OF pi_x ux ux_0] by (simp only: uminus_float.rep_eq of_int_minus cos_minus mult_minus_left) simp also have "\ \ (ub_cos prec (- ?ux))" using lb_cos_minus[OF pi_ux ux_0, of prec] by simp finally have "cos x \ (ub_cos prec (- ?ux))" unfolding cos_periodic_int . } note negative_ux = this { assume "?ux \ ?lpi" and x_ge_0: "0 \ x - k * (2 * pi)" with lpi ux have pi_ux: "?ux \ pi" and ux_0: "0 \ real_of_float ?ux" by simp_all have "(lb_cos prec ?ux) \ cos ?ux" using lb_cos[OF ux_0 pi_ux] by simp also have "\ \ cos (x + (-k) * (2 * pi))" using cos_monotone_0_pi_le[OF x_ge_0 ux pi_ux] by (simp only: of_int_minus cos_minus mult_minus_left) simp finally have "(lb_cos prec ?ux) \ cos x" unfolding cos_periodic_int . } note positive_ux = this show "l \ cos x \ cos x \ u" proof (cases "- ?lpi \ ?lx \ ?ux \ 0") case True with bnds have l: "l = lb_cos prec (-?lx)" and u: "u = ub_cos prec (-?ux)" by (auto simp add: bnds_cos_def Let_def) from True lpi[THEN le_imp_neg_le] lx ux have "- pi \ x - k * (2 * pi)" and "x - k * (2 * pi) \ 0" by auto with True negative_ux negative_lx show ?thesis unfolding l u by simp next case 1: False show ?thesis proof (cases "0 \ ?lx \ ?ux \ ?lpi") case True with bnds 1 have l: "l = lb_cos prec ?ux" and u: "u = ub_cos prec ?lx" by (auto simp add: bnds_cos_def Let_def) from True lpi lx ux have "0 \ x - k * (2 * pi)" and "x - k * (2 * pi) \ pi" by auto with True positive_ux positive_lx show ?thesis unfolding l u by simp next case 2: False show ?thesis proof (cases "- ?lpi \ ?lx \ ?ux \ ?lpi") case Cond: True with bnds 1 2 have l: "l = min (lb_cos prec (-?lx)) (lb_cos prec ?ux)" and u: "u = Float 1 0" by (auto simp add: bnds_cos_def Let_def) show ?thesis unfolding u l using negative_lx positive_ux Cond by (cases "x - k * (2 * pi) < 0") (auto simp add: real_of_float_min) next case 3: False show ?thesis proof (cases "0 \ ?lx \ ?ux \ 2 * ?lpi") case Cond: True with bnds 1 2 3 have l: "l = Float (- 1) 0" and u: "u = max (ub_cos prec ?lx) (ub_cos prec (- (?ux - 2 * ?lpi)))" by (auto simp add: bnds_cos_def Let_def) have "cos x \ real_of_float u" proof (cases "x - k * (2 * pi) < pi") case True hence "x - k * (2 * pi) \ pi" by simp from positive_lx[OF Cond[THEN conjunct1] this] show ?thesis unfolding u by (simp add: real_of_float_max) next case False hence "pi \ x - k * (2 * pi)" by simp hence pi_x: "- pi \ x - k * (2 * pi) - 2 * pi" by simp have "?ux \ 2 * pi" using Cond lpi by auto hence "x - k * (2 * pi) - 2 * pi \ 0" using ux by simp have ux_0: "real_of_float (?ux - 2 * ?lpi) \ 0" using Cond by auto from 2 and Cond have "\ ?ux \ ?lpi" by auto hence "- ?lpi \ ?ux - 2 * ?lpi" by auto hence pi_ux: "- pi \ (?ux - 2 * ?lpi)" using lpi[THEN le_imp_neg_le] by auto have x_le_ux: "x - k * (2 * pi) - 2 * pi \ (?ux - 2 * ?lpi)" using ux lpi by auto have "cos x = cos (x + (-k) * (2 * pi) + (-1::int) * (2 * pi))" unfolding cos_periodic_int .. also have "\ \ cos ((?ux - 2 * ?lpi))" using cos_monotone_minus_pi_0'[OF pi_x x_le_ux ux_0] by (simp only: minus_float.rep_eq of_int_minus of_int_1 mult_minus_left mult_1_left) simp also have "\ = cos ((- (?ux - 2 * ?lpi)))" unfolding uminus_float.rep_eq cos_minus .. also have "\ \ (ub_cos prec (- (?ux - 2 * ?lpi)))" using lb_cos_minus[OF pi_ux ux_0] by simp finally show ?thesis unfolding u by (simp add: real_of_float_max) qed thus ?thesis unfolding l by auto next case 4: False show ?thesis proof (cases "-2 * ?lpi \ ?lx \ ?ux \ 0") case Cond: True with bnds 1 2 3 4 have l: "l = Float (- 1) 0" and u: "u = max (ub_cos prec (?lx + 2 * ?lpi)) (ub_cos prec (-?ux))" by (auto simp add: bnds_cos_def Let_def) have "cos x \ u" proof (cases "-pi < x - k * (2 * pi)") case True hence "-pi \ x - k * (2 * pi)" by simp from negative_ux[OF this Cond[THEN conjunct2]] show ?thesis unfolding u by (simp add: real_of_float_max) next case False hence "x - k * (2 * pi) \ -pi" by simp hence pi_x: "x - k * (2 * pi) + 2 * pi \ pi" by simp have "-2 * pi \ ?lx" using Cond lpi by auto hence "0 \ x - k * (2 * pi) + 2 * pi" using lx by simp have lx_0: "0 \ real_of_float (?lx + 2 * ?lpi)" using Cond lpi by auto from 1 and Cond have "\ -?lpi \ ?lx" by auto hence "?lx + 2 * ?lpi \ ?lpi" by auto hence pi_lx: "(?lx + 2 * ?lpi) \ pi" using lpi[THEN le_imp_neg_le] by auto have lx_le_x: "(?lx + 2 * ?lpi) \ x - k * (2 * pi) + 2 * pi" using lx lpi by auto have "cos x = cos (x + (-k) * (2 * pi) + (1 :: int) * (2 * pi))" unfolding cos_periodic_int .. also have "\ \ cos ((?lx + 2 * ?lpi))" using cos_monotone_0_pi_le[OF lx_0 lx_le_x pi_x] by (simp only: minus_float.rep_eq of_int_minus of_int_1 mult_minus_left mult_1_left) simp also have "\ \ (ub_cos prec (?lx + 2 * ?lpi))" using lb_cos[OF lx_0 pi_lx] by simp finally show ?thesis unfolding u by (simp add: real_of_float_max) qed thus ?thesis unfolding l by auto next case False with bnds 1 2 3 4 show ?thesis by (auto simp add: bnds_cos_def Let_def) qed qed qed qed qed qed lemma bnds_cos_lower: "\x. real_of_float xl \ x \ x \ real_of_float xu \ cos x \ y \ real_of_float (fst (bnds_cos prec xl xu)) \ y" and bnds_cos_upper: "\x. real_of_float xl \ x \ x \ real_of_float xu \ y \ cos x \ y \ real_of_float (snd (bnds_cos prec xl xu))" for xl xu::float and y::real using bnds_cos[of "fst (bnds_cos prec xl xu)" "snd (bnds_cos prec xl xu)" prec] by force+ lift_definition cos_float_interval :: "nat \ float interval \ float interval" is "\prec. \(lx, ux). bnds_cos prec lx ux" using bnds_cos by auto (metis (full_types) order_refl order_trans) lemma lower_cos_float_interval: "lower (cos_float_interval p x) = fst (bnds_cos p (lower x) (upper x))" by transfer auto lemma upper_cos_float_interval: "upper (cos_float_interval p x) = snd (bnds_cos p (lower x) (upper x))" by transfer auto lemma cos_float_interval: "cos ` set_of (real_interval x) \ set_of (real_interval (cos_float_interval p x))" by (auto simp: set_of_eq bnds_cos_lower bnds_cos_upper lower_cos_float_interval upper_cos_float_interval) lemma cos_float_intervalI: "cos x \\<^sub>r cos_float_interval p X" if "x \\<^sub>r X" using cos_float_interval[of X p] that by auto section "Exponential function" subsection "Compute the series of the exponential function" fun ub_exp_horner :: "nat \ nat \ nat \ nat \ float \ float" and lb_exp_horner :: "nat \ nat \ nat \ nat \ float \ float" where "ub_exp_horner prec 0 i k x = 0" | "ub_exp_horner prec (Suc n) i k x = float_plus_up prec (rapprox_rat prec 1 (int k)) (float_round_up prec (x * lb_exp_horner prec n (i + 1) (k * i) x))" | "lb_exp_horner prec 0 i k x = 0" | "lb_exp_horner prec (Suc n) i k x = float_plus_down prec (lapprox_rat prec 1 (int k)) (float_round_down prec (x * ub_exp_horner prec n (i + 1) (k * i) x))" lemma bnds_exp_horner: assumes "real_of_float x \ 0" shows "exp x \ {lb_exp_horner prec (get_even n) 1 1 x .. ub_exp_horner prec (get_odd n) 1 1 x}" proof - have f_eq: "fact (Suc n) = fact n * ((\i::nat. i + 1) ^^ n) 1" for n proof - have F: "\ m. ((\i. i + 1) ^^ n) m = n + m" by (induct n) auto show ?thesis unfolding F by auto qed note bounds = horner_bounds_nonpos[where f="fact" and lb="lb_exp_horner prec" and ub="ub_exp_horner prec" and j'=0 and s=1, OF assms f_eq lb_exp_horner.simps ub_exp_horner.simps] have "lb_exp_horner prec (get_even n) 1 1 x \ exp x" proof - have "lb_exp_horner prec (get_even n) 1 1 x \ (\j = 0.. \ exp x" proof - obtain t where "\t\ \ \real_of_float x\" and "exp x = (\m = 0.. exp t / (fact (get_even n)) * (real_of_float x) ^ (get_even n)" by (auto simp: zero_le_even_power) ultimately show ?thesis using get_odd exp_gt_zero by auto qed finally show ?thesis . qed moreover have "exp x \ ub_exp_horner prec (get_odd n) 1 1 x" proof - have x_less_zero: "real_of_float x ^ get_odd n \ 0" proof (cases "real_of_float x = 0") case True have "(get_odd n) \ 0" using get_odd[THEN odd_pos] by auto thus ?thesis unfolding True power_0_left by auto next case False hence "real_of_float x < 0" using \real_of_float x \ 0\ by auto show ?thesis by (rule less_imp_le, auto simp add: \real_of_float x < 0\) qed obtain t where "\t\ \ \real_of_float x\" and "exp x = (\m = 0.. 0" by (auto intro!: mult_nonneg_nonpos divide_nonpos_pos simp add: x_less_zero) ultimately have "exp x \ (\j = 0.. \ ub_exp_horner prec (get_odd n) 1 1 x" using bounds(2) by auto finally show ?thesis . qed ultimately show ?thesis by auto qed lemma ub_exp_horner_nonneg: "real_of_float x \ 0 \ 0 \ real_of_float (ub_exp_horner prec (get_odd n) (Suc 0) (Suc 0) x)" using bnds_exp_horner[of x prec n] by (intro order_trans[OF exp_ge_zero]) auto subsection "Compute the exponential function on the entire domain" function ub_exp :: "nat \ float \ float" and lb_exp :: "nat \ float \ float" where "lb_exp prec x = (if 0 < x then float_divl prec 1 (ub_exp prec (-x)) else let horner = (\ x. let y = lb_exp_horner prec (get_even (prec + 2)) 1 1 x in if y \ 0 then Float 1 (- 2) else y) in if x < - 1 then power_down_fl prec (horner (float_divl prec x (- floor_fl x))) (nat (- int_floor_fl x)) else horner x)" | "ub_exp prec x = (if 0 < x then float_divr prec 1 (lb_exp prec (-x)) else if x < - 1 then power_up_fl prec (ub_exp_horner prec (get_odd (prec + 2)) 1 1 (float_divr prec x (- floor_fl x))) (nat (- int_floor_fl x)) else ub_exp_horner prec (get_odd (prec + 2)) 1 1 x)" by pat_completeness auto termination by (relation "measure (\ v. let (prec, x) = case_sum id id v in (if 0 < x then 1 else 0))") auto lemma exp_m1_ge_quarter: "(1 / 4 :: real) \ exp (- 1)" proof - have eq4: "4 = Suc (Suc (Suc (Suc 0)))" by auto have "1 / 4 = (Float 1 (- 2))" unfolding Float_num by auto also have "\ \ lb_exp_horner 3 (get_even 3) 1 1 (- 1)" by (subst less_eq_float.rep_eq [symmetric]) code_simp also have "\ \ exp (- 1 :: float)" using bnds_exp_horner[where x="- 1"] by auto finally show ?thesis by simp qed lemma lb_exp_pos: assumes "\ 0 < x" shows "0 < lb_exp prec x" proof - let "?lb_horner x" = "lb_exp_horner prec (get_even (prec + 2)) 1 1 x" let "?horner x" = "let y = ?lb_horner x in if y \ 0 then Float 1 (- 2) else y" have pos_horner: "0 < ?horner x" for x unfolding Let_def by (cases "?lb_horner x \ 0") auto moreover have "0 < real_of_float ((?horner x) ^ num)" for x :: float and num :: nat proof - have "0 < real_of_float (?horner x) ^ num" using \0 < ?horner x\ by simp also have "\ = (?horner x) ^ num" by auto finally show ?thesis . qed ultimately show ?thesis unfolding lb_exp.simps if_not_P[OF \\ 0 < x\] Let_def by (cases "floor_fl x", cases "x < - 1") (auto simp: real_power_up_fl real_power_down_fl intro!: power_up_less power_down_pos) qed lemma exp_boundaries': assumes "x \ 0" shows "exp x \ { (lb_exp prec x) .. (ub_exp prec x)}" proof - let "?lb_exp_horner x" = "lb_exp_horner prec (get_even (prec + 2)) 1 1 x" let "?ub_exp_horner x" = "ub_exp_horner prec (get_odd (prec + 2)) 1 1 x" have "real_of_float x \ 0" and "\ x > 0" using \x \ 0\ by auto show ?thesis proof (cases "x < - 1") case False hence "- 1 \ real_of_float x" by auto show ?thesis proof (cases "?lb_exp_horner x \ 0") case True from \\ x < - 1\ have "- 1 \ real_of_float x" by auto hence "exp (- 1) \ exp x" unfolding exp_le_cancel_iff . from order_trans[OF exp_m1_ge_quarter this] have "Float 1 (- 2) \ exp x" unfolding Float_num . with True show ?thesis using bnds_exp_horner \real_of_float x \ 0\ \\ x > 0\ \\ x < - 1\ by auto next case False thus ?thesis using bnds_exp_horner \real_of_float x \ 0\ \\ x > 0\ \\ x < - 1\ by (auto simp add: Let_def) qed next case True let ?num = "nat (- int_floor_fl x)" have "real_of_int (int_floor_fl x) < - 1" using int_floor_fl[of x] \x < - 1\ by simp hence "real_of_int (int_floor_fl x) < 0" by simp hence "int_floor_fl x < 0" by auto hence "1 \ - int_floor_fl x" by auto hence "0 < nat (- int_floor_fl x)" by auto hence "0 < ?num" by auto hence "real ?num \ 0" by auto have num_eq: "real ?num = - int_floor_fl x" using \0 < nat (- int_floor_fl x)\ by auto have "0 < - int_floor_fl x" using \0 < ?num\[unfolded of_nat_less_iff[symmetric]] by simp hence "real_of_int (int_floor_fl x) < 0" unfolding less_float_def by auto have fl_eq: "real_of_int (- int_floor_fl x) = real_of_float (- floor_fl x)" by (simp add: floor_fl_def int_floor_fl_def) from \0 < - int_floor_fl x\ have "0 \ real_of_float (- floor_fl x)" by (simp add: floor_fl_def int_floor_fl_def) from \real_of_int (int_floor_fl x) < 0\ have "real_of_float (floor_fl x) < 0" by (simp add: floor_fl_def int_floor_fl_def) have "exp x \ ub_exp prec x" proof - have div_less_zero: "real_of_float (float_divr prec x (- floor_fl x)) \ 0" using float_divr_nonpos_pos_upper_bound[OF \real_of_float x \ 0\ \0 \ real_of_float (- floor_fl x)\] unfolding less_eq_float_def zero_float.rep_eq . have "exp x = exp (?num * (x / ?num))" using \real ?num \ 0\ by auto also have "\ = exp (x / ?num) ^ ?num" unfolding exp_of_nat_mult .. also have "\ \ exp (float_divr prec x (- floor_fl x)) ^ ?num" unfolding num_eq fl_eq by (rule power_mono, rule exp_le_cancel_iff[THEN iffD2], rule float_divr) auto also have "\ \ (?ub_exp_horner (float_divr prec x (- floor_fl x))) ^ ?num" unfolding real_of_float_power by (rule power_mono, rule bnds_exp_horner[OF div_less_zero, unfolded atLeastAtMost_iff, THEN conjunct2], auto) also have "\ \ real_of_float (power_up_fl prec (?ub_exp_horner (float_divr prec x (- floor_fl x))) ?num)" by (auto simp add: real_power_up_fl intro!: power_up ub_exp_horner_nonneg div_less_zero) finally show ?thesis unfolding ub_exp.simps if_not_P[OF \\ 0 < x\] if_P[OF \x < - 1\] floor_fl_def Let_def . qed moreover have "lb_exp prec x \ exp x" proof - let ?divl = "float_divl prec x (- floor_fl x)" let ?horner = "?lb_exp_horner ?divl" show ?thesis proof (cases "?horner \ 0") case False hence "0 \ real_of_float ?horner" by auto have div_less_zero: "real_of_float (float_divl prec x (- floor_fl x)) \ 0" using \real_of_float (floor_fl x) < 0\ \real_of_float x \ 0\ by (auto intro!: order_trans[OF float_divl] divide_nonpos_neg) have "(?lb_exp_horner (float_divl prec x (- floor_fl x))) ^ ?num \ exp (float_divl prec x (- floor_fl x)) ^ ?num" using \0 \ real_of_float ?horner\[unfolded floor_fl_def[symmetric]] bnds_exp_horner[OF div_less_zero, unfolded atLeastAtMost_iff, THEN conjunct1] by (auto intro!: power_mono) also have "\ \ exp (x / ?num) ^ ?num" unfolding num_eq fl_eq using float_divl by (auto intro!: power_mono simp del: uminus_float.rep_eq) also have "\ = exp (?num * (x / ?num))" unfolding exp_of_nat_mult .. also have "\ = exp x" using \real ?num \ 0\ by auto finally show ?thesis using False unfolding lb_exp.simps if_not_P[OF \\ 0 < x\] if_P[OF \x < - 1\] int_floor_fl_def Let_def if_not_P[OF False] by (auto simp: real_power_down_fl intro!: power_down_le) next case True have "power_down_fl prec (Float 1 (- 2)) ?num \ (Float 1 (- 2)) ^ ?num" by (metis Float_le_zero_iff less_imp_le linorder_not_less not_numeral_le_zero numeral_One power_down_fl) then have "power_down_fl prec (Float 1 (- 2)) ?num \ real_of_float (Float 1 (- 2)) ^ ?num" by simp also have "real_of_float (floor_fl x) \ 0" and "real_of_float (floor_fl x) \ 0" using \real_of_float (floor_fl x) < 0\ by auto from divide_right_mono_neg[OF floor_fl[of x] \real_of_float (floor_fl x) \ 0\, unfolded divide_self[OF \real_of_float (floor_fl x) \ 0\]] have "- 1 \ x / (- floor_fl x)" unfolding minus_float.rep_eq by auto from order_trans[OF exp_m1_ge_quarter this[unfolded exp_le_cancel_iff[where x="- 1", symmetric]]] have "Float 1 (- 2) \ exp (x / (- floor_fl x))" unfolding Float_num . hence "real_of_float (Float 1 (- 2)) ^ ?num \ exp (x / (- floor_fl x)) ^ ?num" by (metis Float_num(5) power_mono zero_le_divide_1_iff zero_le_numeral) also have "\ = exp x" unfolding num_eq fl_eq exp_of_nat_mult[symmetric] using \real_of_float (floor_fl x) \ 0\ by auto finally show ?thesis unfolding lb_exp.simps if_not_P[OF \\ 0 < x\] if_P[OF \x < - 1\] int_floor_fl_def Let_def if_P[OF True] real_of_float_power . qed qed ultimately show ?thesis by auto qed qed lemma exp_boundaries: "exp x \ { lb_exp prec x .. ub_exp prec x }" proof - show ?thesis proof (cases "0 < x") case False hence "x \ 0" by auto from exp_boundaries'[OF this] show ?thesis . next case True hence "-x \ 0" by auto have "lb_exp prec x \ exp x" proof - from exp_boundaries'[OF \-x \ 0\] have ub_exp: "exp (- real_of_float x) \ ub_exp prec (-x)" unfolding atLeastAtMost_iff minus_float.rep_eq by auto have "float_divl prec 1 (ub_exp prec (-x)) \ 1 / ub_exp prec (-x)" using float_divl[where x=1] by auto also have "\ \ exp x" using ub_exp[unfolded inverse_le_iff_le[OF order_less_le_trans[OF exp_gt_zero ub_exp] exp_gt_zero, symmetric]] unfolding exp_minus nonzero_inverse_inverse_eq[OF exp_not_eq_zero] inverse_eq_divide by auto finally show ?thesis unfolding lb_exp.simps if_P[OF True] . qed moreover have "exp x \ ub_exp prec x" proof - have "\ 0 < -x" using \0 < x\ by auto from exp_boundaries'[OF \-x \ 0\] have lb_exp: "lb_exp prec (-x) \ exp (- real_of_float x)" unfolding atLeastAtMost_iff minus_float.rep_eq by auto have "exp x \ (1 :: float) / lb_exp prec (-x)" using lb_exp lb_exp_pos[OF \\ 0 < -x\, of prec] by (simp del: lb_exp.simps add: exp_minus field_simps) also have "\ \ float_divr prec 1 (lb_exp prec (-x))" using float_divr . finally show ?thesis unfolding ub_exp.simps if_P[OF True] . qed ultimately show ?thesis by auto qed qed lemma bnds_exp: "\(x::real) lx ux. (l, u) = (lb_exp prec lx, ub_exp prec ux) \ x \ {lx .. ux} \ l \ exp x \ exp x \ u" proof (rule allI, rule allI, rule allI, rule impI) fix x :: real and lx ux assume "(l, u) = (lb_exp prec lx, ub_exp prec ux) \ x \ {lx .. ux}" hence l: "lb_exp prec lx = l " and u: "ub_exp prec ux = u" and x: "x \ {lx .. ux}" by auto show "l \ exp x \ exp x \ u" proof show "l \ exp x" proof - from exp_boundaries[of lx prec, unfolded l] have "l \ exp lx" by (auto simp del: lb_exp.simps) also have "\ \ exp x" using x by auto finally show ?thesis . qed show "exp x \ u" proof - have "exp x \ exp ux" using x by auto also have "\ \ u" using exp_boundaries[of ux prec, unfolded u] by (auto simp del: ub_exp.simps) finally show ?thesis . qed qed qed lemmas [simp del] = lb_exp.simps ub_exp.simps lemma lb_exp: "exp x \ y \ lb_exp prec x \ y" and ub_exp: "y \ exp x \ y \ ub_exp prec x" for x::float and y::real using exp_boundaries[of x prec] by auto lift_definition exp_float_interval :: "nat \ float interval \ float interval" is "\prec. \(lx, ux). (lb_exp prec lx, ub_exp prec ux)" by (auto simp: lb_exp ub_exp) lemma lower_exp_float_interval: "lower (exp_float_interval p x) = lb_exp p (lower x)" by transfer auto lemma upper_exp_float_interval: "upper (exp_float_interval p x) = ub_exp p (upper x)" by transfer auto lemma exp_float_interval: "exp ` set_of (real_interval x) \ set_of (real_interval (exp_float_interval p x))" using exp_boundaries by (auto simp: set_of_eq lower_exp_float_interval upper_exp_float_interval lb_exp ub_exp) lemma exp_float_intervalI: "exp x \\<^sub>r exp_float_interval p X" if "x \\<^sub>r X" using exp_float_interval[of X p] that by auto section "Logarithm" subsection "Compute the logarithm series" fun ub_ln_horner :: "nat \ nat \ nat \ float \ float" and lb_ln_horner :: "nat \ nat \ nat \ float \ float" where "ub_ln_horner prec 0 i x = 0" | "ub_ln_horner prec (Suc n) i x = float_plus_up prec (rapprox_rat prec 1 (int i)) (- float_round_down prec (x * lb_ln_horner prec n (Suc i) x))" | "lb_ln_horner prec 0 i x = 0" | "lb_ln_horner prec (Suc n) i x = float_plus_down prec (lapprox_rat prec 1 (int i)) (- float_round_up prec (x * ub_ln_horner prec n (Suc i) x))" lemma ln_bounds: assumes "0 \ x" and "x < 1" shows "(\i=0..<2*n. (- 1) ^ i * (1 / real (i + 1)) * x ^ (Suc i)) \ ln (x + 1)" (is "?lb") and "ln (x + 1) \ (\i=0..<2*n + 1. (- 1) ^ i * (1 / real (i + 1)) * x ^ (Suc i))" (is "?ub") proof - let "?a n" = "(1/real (n +1)) * x ^ (Suc n)" have ln_eq: "(\ i. (- 1) ^ i * ?a i) = ln (x + 1)" using ln_series[of "x + 1"] \0 \ x\ \x < 1\ by auto have "norm x < 1" using assms by auto have "?a \ 0" unfolding Suc_eq_plus1[symmetric] inverse_eq_divide[symmetric] using tendsto_mult[OF LIMSEQ_inverse_real_of_nat LIMSEQ_Suc[OF LIMSEQ_power_zero[OF \norm x < 1\]]] by auto have "0 \ ?a n" for n by (rule mult_nonneg_nonneg) (auto simp: \0 \ x\) have "?a (Suc n) \ ?a n" for n unfolding inverse_eq_divide[symmetric] proof (rule mult_mono) show "0 \ x ^ Suc (Suc n)" by (auto simp add: \0 \ x\) have "x ^ Suc (Suc n) \ x ^ Suc n * 1" unfolding power_Suc2 mult.assoc[symmetric] by (rule mult_left_mono, fact less_imp_le[OF \x < 1\]) (auto simp: \0 \ x\) thus "x ^ Suc (Suc n) \ x ^ Suc n" by auto qed auto from summable_Leibniz'(2,4)[OF \?a \ 0\ \\n. 0 \ ?a n\, OF \\n. ?a (Suc n) \ ?a n\, unfolded ln_eq] show ?lb and ?ub unfolding atLeast0LessThan by auto qed lemma ln_float_bounds: assumes "0 \ real_of_float x" and "real_of_float x < 1" shows "x * lb_ln_horner prec (get_even n) 1 x \ ln (x + 1)" (is "?lb \ ?ln") and "ln (x + 1) \ x * ub_ln_horner prec (get_odd n) 1 x" (is "?ln \ ?ub") proof - obtain ev where ev: "get_even n = 2 * ev" using get_even_double .. obtain od where od: "get_odd n = 2 * od + 1" using get_odd_double .. let "?s n" = "(- 1) ^ n * (1 / real (1 + n)) * (real_of_float x)^(Suc n)" have "?lb \ sum ?s {0 ..< 2 * ev}" unfolding power_Suc2 mult.assoc[symmetric] times_float.rep_eq sum_distrib_right[symmetric] unfolding mult.commute[of "real_of_float x"] ev using horner_bounds(1)[where G="\ i k. Suc k" and F="\x. x" and f="\x. x" and lb="\n i k x. lb_ln_horner prec n k x" and ub="\n i k x. ub_ln_horner prec n k x" and j'=1 and n="2*ev", OF \0 \ real_of_float x\ refl lb_ln_horner.simps ub_ln_horner.simps] \0 \ real_of_float x\ unfolding real_of_float_power by (rule mult_right_mono) also have "\ \ ?ln" using ln_bounds(1)[OF \0 \ real_of_float x\ \real_of_float x < 1\] by auto finally show "?lb \ ?ln" . have "?ln \ sum ?s {0 ..< 2 * od + 1}" using ln_bounds(2)[OF \0 \ real_of_float x\ \real_of_float x < 1\] by auto also have "\ \ ?ub" unfolding power_Suc2 mult.assoc[symmetric] times_float.rep_eq sum_distrib_right[symmetric] unfolding mult.commute[of "real_of_float x"] od using horner_bounds(2)[where G="\ i k. Suc k" and F="\x. x" and f="\x. x" and lb="\n i k x. lb_ln_horner prec n k x" and ub="\n i k x. ub_ln_horner prec n k x" and j'=1 and n="2 * od + 1", OF \0 \ real_of_float x\ refl lb_ln_horner.simps ub_ln_horner.simps] \0 \ real_of_float x\ unfolding real_of_float_power by (rule mult_right_mono) finally show "?ln \ ?ub" . qed lemma ln_add: fixes x :: real assumes "0 < x" and "0 < y" shows "ln (x + y) = ln x + ln (1 + y / x)" proof - have "x \ 0" using assms by auto have "x + y = x * (1 + y / x)" unfolding distrib_left times_divide_eq_right nonzero_mult_div_cancel_left[OF \x \ 0\] by auto moreover have "0 < y / x" using assms by auto hence "0 < 1 + y / x" by auto ultimately show ?thesis using ln_mult assms by auto qed subsection "Compute the logarithm of 2" definition ub_ln2 where "ub_ln2 prec = (let third = rapprox_rat (max prec 1) 1 3 in float_plus_up prec ((Float 1 (- 1) * ub_ln_horner prec (get_odd prec) 1 (Float 1 (- 1)))) (float_round_up prec (third * ub_ln_horner prec (get_odd prec) 1 third)))" definition lb_ln2 where "lb_ln2 prec = (let third = lapprox_rat prec 1 3 in float_plus_down prec ((Float 1 (- 1) * lb_ln_horner prec (get_even prec) 1 (Float 1 (- 1)))) (float_round_down prec (third * lb_ln_horner prec (get_even prec) 1 third)))" lemma ub_ln2: "ln 2 \ ub_ln2 prec" (is "?ub_ln2") and lb_ln2: "lb_ln2 prec \ ln 2" (is "?lb_ln2") proof - let ?uthird = "rapprox_rat (max prec 1) 1 3" let ?lthird = "lapprox_rat prec 1 3" have ln2_sum: "ln 2 = ln (1/2 + 1) + ln (1 / 3 + 1::real)" using ln_add[of "3 / 2" "1 / 2"] by auto have lb3: "?lthird \ 1 / 3" using lapprox_rat[of prec 1 3] by auto hence lb3_ub: "real_of_float ?lthird < 1" by auto have lb3_lb: "0 \ real_of_float ?lthird" using lapprox_rat_nonneg[of 1 3] by auto have ub3: "1 / 3 \ ?uthird" using rapprox_rat[of 1 3] by auto hence ub3_lb: "0 \ real_of_float ?uthird" by auto have lb2: "0 \ real_of_float (Float 1 (- 1))" and ub2: "real_of_float (Float 1 (- 1)) < 1" unfolding Float_num by auto have "0 \ (1::int)" and "0 < (3::int)" by auto have ub3_ub: "real_of_float ?uthird < 1" by (simp add: Float.compute_rapprox_rat Float.compute_lapprox_rat rapprox_posrat_less1) have third_gt0: "(0 :: real) < 1 / 3 + 1" by auto have uthird_gt0: "0 < real_of_float ?uthird + 1" using ub3_lb by auto have lthird_gt0: "0 < real_of_float ?lthird + 1" using lb3_lb by auto show ?ub_ln2 unfolding ub_ln2_def Let_def ln2_sum Float_num(4)[symmetric] proof (rule float_plus_up_le, rule add_mono, fact ln_float_bounds(2)[OF lb2 ub2]) have "ln (1 / 3 + 1) \ ln (real_of_float ?uthird + 1)" unfolding ln_le_cancel_iff[OF third_gt0 uthird_gt0] using ub3 by auto also have "\ \ ?uthird * ub_ln_horner prec (get_odd prec) 1 ?uthird" using ln_float_bounds(2)[OF ub3_lb ub3_ub] . also note float_round_up finally show "ln (1 / 3 + 1) \ float_round_up prec (?uthird * ub_ln_horner prec (get_odd prec) 1 ?uthird)" . qed show ?lb_ln2 unfolding lb_ln2_def Let_def ln2_sum Float_num(4)[symmetric] proof (rule float_plus_down_le, rule add_mono, fact ln_float_bounds(1)[OF lb2 ub2]) have "?lthird * lb_ln_horner prec (get_even prec) 1 ?lthird \ ln (real_of_float ?lthird + 1)" using ln_float_bounds(1)[OF lb3_lb lb3_ub] . note float_round_down_le[OF this] also have "\ \ ln (1 / 3 + 1)" unfolding ln_le_cancel_iff[OF lthird_gt0 third_gt0] using lb3 by auto finally show "float_round_down prec (?lthird * lb_ln_horner prec (get_even prec) 1 ?lthird) \ ln (1 / 3 + 1)" . qed qed subsection "Compute the logarithm in the entire domain" function ub_ln :: "nat \ float \ float option" and lb_ln :: "nat \ float \ float option" where "ub_ln prec x = (if x \ 0 then None else if x < 1 then Some (- the (lb_ln prec (float_divl (max prec 1) 1 x))) else let horner = \x. float_round_up prec (x * ub_ln_horner prec (get_odd prec) 1 x) in if x \ Float 3 (- 1) then Some (horner (x - 1)) else if x < Float 1 1 then Some (float_round_up prec (horner (Float 1 (- 1)) + horner (x * rapprox_rat prec 2 3 - 1))) else let l = bitlen (mantissa x) - 1 in Some (float_plus_up prec (float_round_up prec (ub_ln2 prec * (Float (exponent x + l) 0))) (horner (Float (mantissa x) (- l) - 1))))" | "lb_ln prec x = (if x \ 0 then None else if x < 1 then Some (- the (ub_ln prec (float_divr prec 1 x))) else let horner = \x. float_round_down prec (x * lb_ln_horner prec (get_even prec) 1 x) in if x \ Float 3 (- 1) then Some (horner (x - 1)) else if x < Float 1 1 then Some (float_round_down prec (horner (Float 1 (- 1)) + horner (max (x * lapprox_rat prec 2 3 - 1) 0))) else let l = bitlen (mantissa x) - 1 in Some (float_plus_down prec (float_round_down prec (lb_ln2 prec * (Float (exponent x + l) 0))) (horner (Float (mantissa x) (- l) - 1))))" by pat_completeness auto termination proof (relation "measure (\ v. let (prec, x) = case_sum id id v in (if x < 1 then 1 else 0))", auto) fix prec and x :: float assume "\ real_of_float x \ 0" and "real_of_float x < 1" and "real_of_float (float_divl (max prec (Suc 0)) 1 x) < 1" hence "0 < real_of_float x" "1 \ max prec (Suc 0)" "real_of_float x < 1" by auto from float_divl_pos_less1_bound[OF \0 < real_of_float x\ \real_of_float x < 1\[THEN less_imp_le] \1 \ max prec (Suc 0)\] show False using \real_of_float (float_divl (max prec (Suc 0)) 1 x) < 1\ by auto next fix prec x assume "\ real_of_float x \ 0" and "real_of_float x < 1" and "real_of_float (float_divr prec 1 x) < 1" hence "0 < x" by auto from float_divr_pos_less1_lower_bound[OF \0 < x\, of prec] \real_of_float x < 1\ show False using \real_of_float (float_divr prec 1 x) < 1\ by auto qed lemmas float_pos_eq_mantissa_pos = mantissa_pos_iff[symmetric] lemma Float_pos_eq_mantissa_pos: "Float m e > 0 \ m > 0" using powr_gt_zero[of 2 "e"] by (auto simp add: zero_less_mult_iff zero_float_def simp del: powr_gt_zero dest: less_zeroE) lemma Float_representation_aux: fixes m e defines [THEN meta_eq_to_obj_eq]: "x \ Float m e" assumes "x > 0" shows "Float (exponent x + (bitlen (mantissa x) - 1)) 0 = Float (e + (bitlen m - 1)) 0" (is ?th1) and "Float (mantissa x) (- (bitlen (mantissa x) - 1)) = Float m ( - (bitlen m - 1))" (is ?th2) proof - from assms have mantissa_pos: "m > 0" "mantissa x > 0" using Float_pos_eq_mantissa_pos[of m e] float_pos_eq_mantissa_pos[of x] by simp_all thus ?th1 using bitlen_Float[of m e] assms by (auto simp add: zero_less_mult_iff intro!: arg_cong2[where f=Float]) have "x \ 0" unfolding zero_float_def[symmetric] using \0 < x\ by auto from denormalize_shift[OF x_def this] obtain i where i: "m = mantissa x * 2 ^ i" "e = exponent x - int i" . have "2 powr (1 - (real_of_int (bitlen (mantissa x)) + real_of_int i)) = 2 powr (1 - (real_of_int (bitlen (mantissa x)))) * inverse (2 powr (real i))" by (simp add: powr_minus[symmetric] powr_add[symmetric] field_simps) hence "real_of_int (mantissa x) * 2 powr (1 - real_of_int (bitlen (mantissa x))) = (real_of_int (mantissa x) * 2 ^ i) * 2 powr (1 - real_of_int (bitlen (mantissa x * 2 ^ i)))" using \mantissa x > 0\ by (simp add: powr_realpow) then show ?th2 unfolding i by (auto simp: real_of_float_eq) qed lemma compute_ln[code]: fixes m e defines "x \ Float m e" shows "ub_ln prec x = (if x \ 0 then None else if x < 1 then Some (- the (lb_ln prec (float_divl (max prec 1) 1 x))) else let horner = \x. float_round_up prec (x * ub_ln_horner prec (get_odd prec) 1 x) in if x \ Float 3 (- 1) then Some (horner (x - 1)) else if x < Float 1 1 then Some (float_round_up prec (horner (Float 1 (- 1)) + horner (x * rapprox_rat prec 2 3 - 1))) else let l = bitlen m - 1 in Some (float_plus_up prec (float_round_up prec (ub_ln2 prec * (Float (e + l) 0))) (horner (Float m (- l) - 1))))" (is ?th1) and "lb_ln prec x = (if x \ 0 then None else if x < 1 then Some (- the (ub_ln prec (float_divr prec 1 x))) else let horner = \x. float_round_down prec (x * lb_ln_horner prec (get_even prec) 1 x) in if x \ Float 3 (- 1) then Some (horner (x - 1)) else if x < Float 1 1 then Some (float_round_down prec (horner (Float 1 (- 1)) + horner (max (x * lapprox_rat prec 2 3 - 1) 0))) else let l = bitlen m - 1 in Some (float_plus_down prec (float_round_down prec (lb_ln2 prec * (Float (e + l) 0))) (horner (Float m (- l) - 1))))" (is ?th2) proof - from assms Float_pos_eq_mantissa_pos have "x > 0 \ m > 0" by simp thus ?th1 ?th2 using Float_representation_aux[of m e] unfolding x_def[symmetric] by (auto dest: not_le_imp_less) qed lemma ln_shifted_float: assumes "0 < m" shows "ln (Float m e) = ln 2 * (e + (bitlen m - 1)) + ln (Float m (- (bitlen m - 1)))" proof - let ?B = "2^nat (bitlen m - 1)" define bl where "bl = bitlen m - 1" have "0 < real_of_int m" and "\X. (0 :: real) < 2^X" and "0 < (2 :: real)" and "m \ 0" using assms by auto hence "0 \ bl" by (simp add: bitlen_alt_def bl_def) show ?thesis proof (cases "0 \ e") case True thus ?thesis unfolding bl_def[symmetric] using \0 < real_of_int m\ \0 \ bl\ apply (simp add: ln_mult) apply (cases "e=0") apply (cases "bl = 0", simp_all add: powr_minus ln_inverse ln_powr) apply (cases "bl = 0", simp_all add: powr_minus ln_inverse ln_powr field_simps) done next case False hence "0 < -e" by auto have lne: "ln (2 powr real_of_int e) = ln (inverse (2 powr - e))" by (simp add: powr_minus) hence pow_gt0: "(0::real) < 2^nat (-e)" by auto hence inv_gt0: "(0::real) < inverse (2^nat (-e))" by auto show ?thesis using False unfolding bl_def[symmetric] using \0 < real_of_int m\ \0 \ bl\ by (auto simp add: lne ln_mult ln_powr ln_div field_simps) qed qed lemma ub_ln_lb_ln_bounds': assumes "1 \ x" shows "the (lb_ln prec x) \ ln x \ ln x \ the (ub_ln prec x)" (is "?lb \ ?ln \ ?ln \ ?ub") proof (cases "x < Float 1 1") case True hence "real_of_float (x - 1) < 1" and "real_of_float x < 2" by auto have "\ x \ 0" and "\ x < 1" using \1 \ x\ by auto hence "0 \ real_of_float (x - 1)" using \1 \ x\ by auto have [simp]: "(Float 3 (- 1)) = 3 / 2" by simp show ?thesis proof (cases "x \ Float 3 (- 1)") case True show ?thesis unfolding lb_ln.simps unfolding ub_ln.simps Let_def using ln_float_bounds[OF \0 \ real_of_float (x - 1)\ \real_of_float (x - 1) < 1\, of prec] \\ x \ 0\ \\ x < 1\ True by (auto intro!: float_round_down_le float_round_up_le) next case False hence *: "3 / 2 < x" by auto with ln_add[of "3 / 2" "x - 3 / 2"] have add: "ln x = ln (3 / 2) + ln (real_of_float x * 2 / 3)" by (auto simp add: algebra_simps diff_divide_distrib) let "?ub_horner x" = "float_round_up prec (x * ub_ln_horner prec (get_odd prec) 1 x)" let "?lb_horner x" = "float_round_down prec (x * lb_ln_horner prec (get_even prec) 1 x)" { have up: "real_of_float (rapprox_rat prec 2 3) \ 1" by (rule rapprox_rat_le1) simp_all have low: "2 / 3 \ rapprox_rat prec 2 3" by (rule order_trans[OF _ rapprox_rat]) simp from mult_less_le_imp_less[OF * low] * have pos: "0 < real_of_float (x * rapprox_rat prec 2 3 - 1)" by auto have "ln (real_of_float x * 2/3) \ ln (real_of_float (x * rapprox_rat prec 2 3 - 1) + 1)" proof (rule ln_le_cancel_iff[symmetric, THEN iffD1]) show "real_of_float x * 2 / 3 \ real_of_float (x * rapprox_rat prec 2 3 - 1) + 1" using * low by auto show "0 < real_of_float x * 2 / 3" using * by simp show "0 < real_of_float (x * rapprox_rat prec 2 3 - 1) + 1" using pos by auto qed also have "\ \ ?ub_horner (x * rapprox_rat prec 2 3 - 1)" proof (rule float_round_up_le, rule ln_float_bounds(2)) from mult_less_le_imp_less[OF \real_of_float x < 2\ up] low * show "real_of_float (x * rapprox_rat prec 2 3 - 1) < 1" by auto show "0 \ real_of_float (x * rapprox_rat prec 2 3 - 1)" using pos by auto qed finally have "ln x \ ?ub_horner (Float 1 (-1)) + ?ub_horner ((x * rapprox_rat prec 2 3 - 1))" using ln_float_bounds(2)[of "Float 1 (- 1)" prec prec] add by (auto intro!: add_mono float_round_up_le) note float_round_up_le[OF this, of prec] } moreover { let ?max = "max (x * lapprox_rat prec 2 3 - 1) 0" have up: "lapprox_rat prec 2 3 \ 2/3" by (rule order_trans[OF lapprox_rat], simp) have low: "0 \ real_of_float (lapprox_rat prec 2 3)" using lapprox_rat_nonneg[of 2 3 prec] by simp have "?lb_horner ?max \ ln (real_of_float ?max + 1)" proof (rule float_round_down_le, rule ln_float_bounds(1)) from mult_less_le_imp_less[OF \real_of_float x < 2\ up] * low show "real_of_float ?max < 1" by (cases "real_of_float (lapprox_rat prec 2 3) = 0", auto simp add: real_of_float_max) show "0 \ real_of_float ?max" by (auto simp add: real_of_float_max) qed also have "\ \ ln (real_of_float x * 2/3)" proof (rule ln_le_cancel_iff[symmetric, THEN iffD1]) show "0 < real_of_float ?max + 1" by (auto simp add: real_of_float_max) show "0 < real_of_float x * 2/3" using * by auto show "real_of_float ?max + 1 \ real_of_float x * 2/3" using * up by (cases "0 < real_of_float x * real_of_float (lapprox_posrat prec 2 3) - 1", auto simp add: max_def) qed finally have "?lb_horner (Float 1 (- 1)) + ?lb_horner ?max \ ln x" using ln_float_bounds(1)[of "Float 1 (- 1)" prec prec] add by (auto intro!: add_mono float_round_down_le) note float_round_down_le[OF this, of prec] } ultimately show ?thesis unfolding lb_ln.simps unfolding ub_ln.simps Let_def using \\ x \ 0\ \\ x < 1\ True False by auto qed next case False hence "\ x \ 0" and "\ x < 1" "0 < x" "\ x \ Float 3 (- 1)" using \1 \ x\ by auto show ?thesis proof - define m where "m = mantissa x" define e where "e = exponent x" from Float_mantissa_exponent[of x] have Float: "x = Float m e" by (simp add: m_def e_def) let ?s = "Float (e + (bitlen m - 1)) 0" let ?x = "Float m (- (bitlen m - 1))" have "0 < m" and "m \ 0" using \0 < x\ Float powr_gt_zero[of 2 e] by (auto simp add: zero_less_mult_iff) define bl where "bl = bitlen m - 1" then have bitlen: "bitlen m = bl + 1" by simp hence "bl \ 0" using \m > 0\ by (auto simp add: bitlen_alt_def) have "1 \ Float m e" using \1 \ x\ Float unfolding less_eq_float_def by auto from bitlen_div[OF \0 < m\] float_gt1_scale[OF \1 \ Float m e\] \bl \ 0\ have x_bnds: "0 \ real_of_float (?x - 1)" "real_of_float (?x - 1) < 1" using abs_real_le_2_powr_bitlen [of m] \m > 0\ by (simp_all add: bitlen powr_realpow [symmetric] powr_minus powr_add field_simps) { have "float_round_down prec (lb_ln2 prec * ?s) \ ln 2 * (e + (bitlen m - 1))" (is "real_of_float ?lb2 \ _") apply (rule float_round_down_le) unfolding nat_0 power_0 mult_1_right times_float.rep_eq using lb_ln2[of prec] proof (rule mult_mono) from float_gt1_scale[OF \1 \ Float m e\] show "0 \ real_of_float (Float (e + (bitlen m - 1)) 0)" by simp qed auto moreover from ln_float_bounds(1)[OF x_bnds] have "float_round_down prec ((?x - 1) * lb_ln_horner prec (get_even prec) 1 (?x - 1)) \ ln ?x" (is "real_of_float ?lb_horner \ _") by (auto intro!: float_round_down_le) ultimately have "float_plus_down prec ?lb2 ?lb_horner \ ln x" unfolding Float ln_shifted_float[OF \0 < m\, of e] by (auto intro!: float_plus_down_le) } moreover { from ln_float_bounds(2)[OF x_bnds] have "ln ?x \ float_round_up prec ((?x - 1) * ub_ln_horner prec (get_odd prec) 1 (?x - 1))" (is "_ \ real_of_float ?ub_horner") by (auto intro!: float_round_up_le) moreover have "ln 2 * (e + (bitlen m - 1)) \ float_round_up prec (ub_ln2 prec * ?s)" (is "_ \ real_of_float ?ub2") apply (rule float_round_up_le) unfolding nat_0 power_0 mult_1_right times_float.rep_eq using ub_ln2[of prec] proof (rule mult_mono) from float_gt1_scale[OF \1 \ Float m e\] show "0 \ real_of_int (e + (bitlen m - 1))" by auto have "0 \ ln (2 :: real)" by simp thus "0 \ real_of_float (ub_ln2 prec)" using ub_ln2[of prec] by arith qed auto ultimately have "ln x \ float_plus_up prec ?ub2 ?ub_horner" unfolding Float ln_shifted_float[OF \0 < m\, of e] by (auto intro!: float_plus_up_le) } ultimately show ?thesis unfolding lb_ln.simps unfolding ub_ln.simps unfolding if_not_P[OF \\ x \ 0\] if_not_P[OF \\ x < 1\] if_not_P[OF False] if_not_P[OF \\ x \ Float 3 (- 1)\] Let_def unfolding plus_float.rep_eq e_def[symmetric] m_def[symmetric] by simp qed qed lemma ub_ln_lb_ln_bounds: assumes "0 < x" shows "the (lb_ln prec x) \ ln x \ ln x \ the (ub_ln prec x)" (is "?lb \ ?ln \ ?ln \ ?ub") proof (cases "x < 1") case False hence "1 \ x" unfolding less_float_def less_eq_float_def by auto show ?thesis using ub_ln_lb_ln_bounds'[OF \1 \ x\] . next case True have "\ x \ 0" using \0 < x\ by auto from True have "real_of_float x \ 1" "x \ 1" by simp_all have "0 < real_of_float x" and "real_of_float x \ 0" using \0 < x\ by auto hence A: "0 < 1 / real_of_float x" by auto { let ?divl = "float_divl (max prec 1) 1 x" have A': "1 \ ?divl" using float_divl_pos_less1_bound[OF \0 < real_of_float x\ \real_of_float x \ 1\] by auto hence B: "0 < real_of_float ?divl" by auto have "ln ?divl \ ln (1 / x)" unfolding ln_le_cancel_iff[OF B A] using float_divl[of _ 1 x] by auto hence "ln x \ - ln ?divl" unfolding nonzero_inverse_eq_divide[OF \real_of_float x \ 0\, symmetric] ln_inverse[OF \0 < real_of_float x\] by auto from this ub_ln_lb_ln_bounds'[OF A', THEN conjunct1, THEN le_imp_neg_le] have "?ln \ - the (lb_ln prec ?divl)" unfolding uminus_float.rep_eq by (rule order_trans) } moreover { let ?divr = "float_divr prec 1 x" have A': "1 \ ?divr" using float_divr_pos_less1_lower_bound[OF \0 < x\ \x \ 1\] unfolding less_eq_float_def less_float_def by auto hence B: "0 < real_of_float ?divr" by auto have "ln (1 / x) \ ln ?divr" unfolding ln_le_cancel_iff[OF A B] using float_divr[of 1 x] by auto hence "- ln ?divr \ ln x" unfolding nonzero_inverse_eq_divide[OF \real_of_float x \ 0\, symmetric] ln_inverse[OF \0 < real_of_float x\] by auto from ub_ln_lb_ln_bounds'[OF A', THEN conjunct2, THEN le_imp_neg_le] this have "- the (ub_ln prec ?divr) \ ?ln" unfolding uminus_float.rep_eq by (rule order_trans) } ultimately show ?thesis unfolding lb_ln.simps[where x=x] ub_ln.simps[where x=x] unfolding if_not_P[OF \\ x \ 0\] if_P[OF True] by auto qed lemma lb_ln: assumes "Some y = lb_ln prec x" shows "y \ ln x" and "0 < real_of_float x" proof - have "0 < x" proof (rule ccontr) assume "\ 0 < x" hence "x \ 0" unfolding less_eq_float_def less_float_def by auto thus False using assms by auto qed thus "0 < real_of_float x" by auto have "the (lb_ln prec x) \ ln x" using ub_ln_lb_ln_bounds[OF \0 < x\] .. thus "y \ ln x" unfolding assms[symmetric] by auto qed lemma ub_ln: assumes "Some y = ub_ln prec x" shows "ln x \ y" and "0 < real_of_float x" proof - have "0 < x" proof (rule ccontr) assume "\ 0 < x" hence "x \ 0" by auto thus False using assms by auto qed thus "0 < real_of_float x" by auto have "ln x \ the (ub_ln prec x)" using ub_ln_lb_ln_bounds[OF \0 < x\] .. thus "ln x \ y" unfolding assms[symmetric] by auto qed lemma bnds_ln: "\(x::real) lx ux. (Some l, Some u) = (lb_ln prec lx, ub_ln prec ux) \ x \ {lx .. ux} \ l \ ln x \ ln x \ u" proof (rule allI, rule allI, rule allI, rule impI) fix x :: real fix lx ux assume "(Some l, Some u) = (lb_ln prec lx, ub_ln prec ux) \ x \ {lx .. ux}" hence l: "Some l = lb_ln prec lx " and u: "Some u = ub_ln prec ux" and x: "x \ {lx .. ux}" by auto have "ln ux \ u" and "0 < real_of_float ux" using ub_ln u by auto have "l \ ln lx" and "0 < real_of_float lx" and "0 < x" using lb_ln[OF l] x by auto from ln_le_cancel_iff[OF \0 < real_of_float lx\ \0 < x\] \l \ ln lx\ have "l \ ln x" using x unfolding atLeastAtMost_iff by auto moreover from ln_le_cancel_iff[OF \0 < x\ \0 < real_of_float ux\] \ln ux \ real_of_float u\ have "ln x \ u" using x unfolding atLeastAtMost_iff by auto ultimately show "l \ ln x \ ln x \ u" .. qed lemmas [simp del] = lb_ln.simps ub_ln.simps lemma lb_lnD: "y \ ln x \ 0 < real_of_float x" if "lb_ln prec x = Some y" using lb_ln[OF that[symmetric]] by auto lemma ub_lnD: "ln x \ y\ 0 < real_of_float x" if "ub_ln prec x = Some y" using ub_ln[OF that[symmetric]] by auto lift_definition(code_dt) ln_float_interval :: "nat \ float interval \ float interval option" is "\prec. \(lx, ux). Option.bind (lb_ln prec lx) (\l. Option.bind (ub_ln prec ux) (\u. Some (l, u)))" by (auto simp: pred_option_def bind_eq_Some_conv ln_le_cancel_iff[symmetric] simp del: ln_le_cancel_iff dest!: lb_lnD ub_lnD) lemma ln_float_interval_eq_Some_conv: "ln_float_interval p x = Some y \ lb_ln p (lower x) = Some (lower y) \ ub_ln p (upper x) = Some (upper y)" by transfer (auto simp: bind_eq_Some_conv) lemma ln_float_interval: "ln ` set_of (real_interval x) \ set_of (real_interval y)" if "ln_float_interval p x = Some y" using that lb_ln[of "lower y" p "lower x"] ub_ln[of "lower y" p "lower x"] apply (auto simp add: set_of_eq ln_float_interval_eq_Some_conv ln_le_cancel_iff) apply (meson less_le_trans ln_less_cancel_iff not_le) by (meson less_le_trans ln_less_cancel_iff not_le ub_lnD) lemma ln_float_intervalI: "ln x \ set_of' (ln_float_interval p X)" if "x \\<^sub>r X" using ln_float_interval[of p X] that by (auto simp: set_of'_def split: option.splits) lemma ln_float_interval_eqI: "ln x \\<^sub>r IVL" if "ln_float_interval p X = Some IVL" "x \\<^sub>r X" using ln_float_intervalI[of x X p] that by (auto simp: set_of'_def split: option.splits) section \Real power function\ definition bnds_powr :: "nat \ float \ float \ float \ float \ (float \ float) option" where "bnds_powr prec l1 u1 l2 u2 = ( if l1 = 0 \ u1 = 0 then Some (0, 0) else if l1 = 0 \ l2 \ 1 then let uln = the (ub_ln prec u1) in Some (0, ub_exp prec (float_round_up prec (uln * (if uln \ 0 then u2 else l2)))) else if l1 \ 0 then None else Some (map_bnds lb_exp ub_exp prec (bnds_mult prec (the (lb_ln prec l1)) (the (ub_ln prec u1)) l2 u2)))" lemma mono_exp_real: "mono (exp :: real \ real)" by (auto simp: mono_def) lemma ub_exp_nonneg: "real_of_float (ub_exp prec x) \ 0" proof - have "0 \ exp (real_of_float x)" by simp also from exp_boundaries[of x prec] have "\ \ real_of_float (ub_exp prec x)" by simp finally show ?thesis . qed lemma bnds_powr: assumes lu: "Some (l, u) = bnds_powr prec l1 u1 l2 u2" assumes x: "x \ {real_of_float l1..real_of_float u1}" assumes y: "y \ {real_of_float l2..real_of_float u2}" shows "x powr y \ {real_of_float l..real_of_float u}" proof - consider "l1 = 0" "u1 = 0" | "l1 = 0" "u1 \ 0" "l2 \ 1" | "l1 \ 0" "\(l1 = 0 \ (u1 = 0 \ l2 \ 1))" | "l1 > 0" by force thus ?thesis proof cases assume "l1 = 0" "u1 = 0" with x lu show ?thesis by (auto simp: bnds_powr_def) next assume A: "l1 = 0" "u1 \ 0" "l2 \ 1" define uln where "uln = the (ub_ln prec u1)" show ?thesis proof (cases "x = 0") case False with A x y have "x powr y = exp (ln x * y)" by (simp add: powr_def) also { from A x False have "ln x \ ln (real_of_float u1)" by simp also from ub_ln_lb_ln_bounds[of u1 prec] A y x False have "ln (real_of_float u1) \ real_of_float uln" by (simp add: uln_def del: lb_ln.simps) also from A x y have "\ * y \ real_of_float uln * (if uln \ 0 then u2 else l2)" by (auto intro: mult_left_mono mult_left_mono_neg) also have "\ \ real_of_float (float_round_up prec (uln * (if uln \ 0 then u2 else l2)))" by (simp add: float_round_up_le) finally have "ln x * y \ \" using A y by - simp } also have "exp (real_of_float (float_round_up prec (uln * (if uln \ 0 then u2 else l2)))) \ real_of_float (ub_exp prec (float_round_up prec (uln * (if uln \ 0 then u2 else l2))))" using exp_boundaries by simp finally show ?thesis using A x y lu by (simp add: bnds_powr_def uln_def Let_def del: lb_ln.simps ub_ln.simps) qed (insert x y lu A, simp_all add: bnds_powr_def Let_def ub_exp_nonneg del: lb_ln.simps ub_ln.simps) next assume "l1 \ 0" "\(l1 = 0 \ (u1 = 0 \ l2 \ 1))" with lu show ?thesis by (simp add: bnds_powr_def split: if_split_asm) next assume l1: "l1 > 0" obtain lm um where lmum: "(lm, um) = bnds_mult prec (the (lb_ln prec l1)) (the (ub_ln prec u1)) l2 u2" by (cases "bnds_mult prec (the (lb_ln prec l1)) (the (ub_ln prec u1)) l2 u2") simp with l1 have "(l, u) = map_bnds lb_exp ub_exp prec (lm, um)" using lu by (simp add: bnds_powr_def del: lb_ln.simps ub_ln.simps split: if_split_asm) hence "exp (ln x * y) \ {real_of_float l..real_of_float u}" proof (rule map_bnds[OF _ mono_exp_real], goal_cases) case 1 let ?lln = "the (lb_ln prec l1)" and ?uln = "the (ub_ln prec u1)" from ub_ln_lb_ln_bounds[of l1 prec] ub_ln_lb_ln_bounds[of u1 prec] x l1 have "real_of_float ?lln \ ln (real_of_float l1) \ ln (real_of_float u1) \ real_of_float ?uln" by (auto simp del: lb_ln.simps ub_ln.simps) moreover from l1 x have "ln (real_of_float l1) \ ln x \ ln x \ ln (real_of_float u1)" by auto ultimately have ln: "real_of_float ?lln \ ln x \ ln x \ real_of_float ?uln" by simp from lmum show ?case by (rule bnds_mult) (insert y ln, simp_all) qed (insert exp_boundaries[of lm prec] exp_boundaries[of um prec], simp_all) with x l1 show ?thesis by (simp add: powr_def mult_ac) qed qed lift_definition(code_dt) powr_float_interval :: "nat \ float interval \ float interval \ float interval option" is "\prec. \(l1, u1). \(l2, u2). bnds_powr prec l1 u1 l2 u2" by (auto simp: pred_option_def dest!: bnds_powr[OF sym]) lemma powr_float_interval: "{x powr y | x y. x \ set_of (real_interval X) \ y \ set_of (real_interval Y)} \ set_of (real_interval R)" if "powr_float_interval prec X Y = Some R" using that by transfer (auto dest!: bnds_powr[OF sym]) lemma powr_float_intervalI: "x powr y \ set_of' (powr_float_interval p X Y)" if "x \\<^sub>r X" "y \\<^sub>r Y" using powr_float_interval[of p X Y] that by (auto simp: set_of'_def split: option.splits) lemma powr_float_interval_eqI: "x powr y \\<^sub>r IVL" if "powr_float_interval p X Y = Some IVL" "x \\<^sub>r X" "y \\<^sub>r Y" using powr_float_intervalI[of x X y Y p] that by (auto simp: set_of'_def split: option.splits) end end \ No newline at end of file diff --git a/src/HOL/Library/Interval.thy b/src/HOL/Library/Interval.thy --- a/src/HOL/Library/Interval.thy +++ b/src/HOL/Library/Interval.thy @@ -1,860 +1,860 @@ (* Title: Interval Author: Christoph Traut, TU Muenchen Fabian Immler, TU Muenchen *) section \Interval Type\ theory Interval imports Complex_Main Lattice_Algebras Set_Algebras begin text \A type of non-empty, closed intervals.\ typedef (overloaded) 'a interval = "{(a::'a::preorder, b). a \ b}" morphisms bounds_of_interval Interval by auto setup_lifting type_definition_interval lift_definition lower::"('a::preorder) interval \ 'a" is fst . lift_definition upper::"('a::preorder) interval \ 'a" is snd . lemma interval_eq_iff: "a = b \ lower a = lower b \ upper a = upper b" by transfer auto lemma interval_eqI: "lower a = lower b \ upper a = upper b \ a = b" by (auto simp: interval_eq_iff) lemma lower_le_upper[simp]: "lower i \ upper i" by transfer auto lift_definition set_of :: "'a::preorder interval \ 'a set" is "\x. {fst x .. snd x}" . lemma set_of_eq: "set_of x = {lower x .. upper x}" by transfer simp context notes [[typedef_overloaded]] begin lift_definition(code_dt) Interval'::"'a::preorder \ 'a::preorder \ 'a interval option" is "\a b. if a \ b then Some (a, b) else None" by auto lemma Interval'_split: "P (Interval' a b) \ (\ivl. a \ b \ lower ivl = a \ upper ivl = b \ P (Some ivl)) \ (\a\b \ P None)" by transfer auto lemma Interval'_split_asm: "P (Interval' a b) \ \((\ivl. a \ b \ lower ivl = a \ upper ivl = b \ \P (Some ivl)) \ (\a\b \ \P None))" unfolding Interval'_split by auto lemmas Interval'_splits = Interval'_split Interval'_split_asm lemma Interval'_eq_Some: "Interval' a b = Some i \ lower i = a \ upper i = b" by (simp split: Interval'_splits) end instantiation "interval" :: ("{preorder,equal}") equal begin definition "equal_class.equal a b \ (lower a = lower b) \ (upper a = upper b)" instance proof qed (simp add: equal_interval_def interval_eq_iff) end instantiation interval :: ("preorder") ord begin definition less_eq_interval :: "'a interval \ 'a interval \ bool" where "less_eq_interval a b \ lower b \ lower a \ upper a \ upper b" definition less_interval :: "'a interval \ 'a interval \ bool" where "less_interval x y = (x \ y \ \ y \ x)" instance proof qed end instantiation interval :: ("lattice") semilattice_sup begin lift_definition sup_interval :: "'a interval \ 'a interval \ 'a interval" is "\(a, b) (c, d). (inf a c, sup b d)" by (auto simp: le_infI1 le_supI1) lemma lower_sup[simp]: "lower (sup A B) = inf (lower A) (lower B)" by transfer auto lemma upper_sup[simp]: "upper (sup A B) = sup (upper A) (upper B)" by transfer auto instance proof qed (auto simp: less_eq_interval_def less_interval_def interval_eq_iff) end lemma set_of_interval_union: "set_of A \ set_of B \ set_of (sup A B)" for A::"'a::lattice interval" by (auto simp: set_of_eq) lemma interval_union_commute: "sup A B = sup B A" for A::"'a::lattice interval" by (auto simp add: interval_eq_iff inf.commute sup.commute) lemma interval_union_mono1: "set_of a \ set_of (sup a A)" for A :: "'a::lattice interval" using set_of_interval_union by blast lemma interval_union_mono2: "set_of A \ set_of (sup a A)" for A :: "'a::lattice interval" using set_of_interval_union by blast lift_definition interval_of :: "'a::preorder \ 'a interval" is "\x. (x, x)" by auto lemma lower_interval_of[simp]: "lower (interval_of a) = a" by transfer auto lemma upper_interval_of[simp]: "upper (interval_of a) = a" by transfer auto definition width :: "'a::{preorder,minus} interval \ 'a" where "width i = upper i - lower i" instantiation "interval" :: ("ordered_ab_semigroup_add") ab_semigroup_add begin lift_definition plus_interval::"'a interval \ 'a interval \ 'a interval" is "\(a, b). \(c, d). (a + c, b + d)" by (auto intro!: add_mono) lemma lower_plus[simp]: "lower (plus A B) = plus (lower A) (lower B)" by transfer auto lemma upper_plus[simp]: "upper (plus A B) = plus (upper A) (upper B)" by transfer auto instance proof qed (auto simp: interval_eq_iff less_eq_interval_def ac_simps) end instance "interval" :: ("{ordered_ab_semigroup_add, lattice}") ordered_ab_semigroup_add proof qed (auto simp: less_eq_interval_def intro!: add_mono) instantiation "interval" :: ("{preorder,zero}") zero begin lift_definition zero_interval::"'a interval" is "(0, 0)" by auto lemma lower_zero[simp]: "lower 0 = 0" by transfer auto lemma upper_zero[simp]: "upper 0 = 0" by transfer auto instance proof qed end instance "interval" :: ("{ordered_comm_monoid_add}") comm_monoid_add proof qed (auto simp: interval_eq_iff) instance "interval" :: ("{ordered_comm_monoid_add,lattice}") ordered_comm_monoid_add .. instantiation "interval" :: ("{ordered_ab_group_add}") uminus begin lift_definition uminus_interval::"'a interval \ 'a interval" is "\(a, b). (-b, -a)" by auto lemma lower_uminus[simp]: "lower (- A) = - upper A" by transfer auto lemma upper_uminus[simp]: "upper (- A) = - lower A" by transfer auto instance .. end instantiation "interval" :: ("{ordered_ab_group_add}") minus begin definition minus_interval::"'a interval \ 'a interval \ 'a interval" where "minus_interval a b = a + - b" lemma lower_minus[simp]: "lower (minus A B) = minus (lower A) (upper B)" by (auto simp: minus_interval_def) lemma upper_minus[simp]: "upper (minus A B) = minus (upper A) (lower B)" by (auto simp: minus_interval_def) instance .. end instantiation "interval" :: (linordered_semiring) times begin lift_definition times_interval :: "'a interval \ 'a interval \ 'a interval" is "\(a1, a2). \(b1, b2). (let x1 = a1 * b1; x2 = a1 * b2; x3 = a2 * b1; x4 = a2 * b2 in (min x1 (min x2 (min x3 x4)), max x1 (max x2 (max x3 x4))))" by (auto simp: Let_def intro!: min.coboundedI1 max.coboundedI1) lemma lower_times: "lower (times A B) = Min {lower A * lower B, lower A * upper B, upper A * lower B, upper A * upper B}" by transfer (auto simp: Let_def) lemma upper_times: "upper (times A B) = Max {lower A * lower B, lower A * upper B, upper A * lower B, upper A * upper B}" by transfer (auto simp: Let_def) instance .. end lemma interval_eq_set_of_iff: "X = Y \ set_of X = set_of Y" for X Y::"'a::order interval" by (auto simp: set_of_eq interval_eq_iff) subsection \Membership\ abbreviation (in preorder) in_interval ("(_/ \\<^sub>i _)" [51, 51] 50) where "in_interval x X \ x \ set_of X" lemma in_interval_to_interval[intro!]: "a \\<^sub>i interval_of a" by (auto simp: set_of_eq) lemma plus_in_intervalI: fixes x y :: "'a :: ordered_ab_semigroup_add" shows "x \\<^sub>i X \ y \\<^sub>i Y \ x + y \\<^sub>i X + Y" by (simp add: add_mono_thms_linordered_semiring(1) set_of_eq) lemma connected_set_of[intro, simp]: "connected (set_of X)" for X::"'a::linear_continuum_topology interval" by (auto simp: set_of_eq ) lemma ex_sum_in_interval_lemma: "\xa\{la .. ua}. \xb\{lb .. ub}. x = xa + xb" if "la \ ua" "lb \ ub" "la + lb \ x" "x \ ua + ub" "ua - la \ ub - lb" for la b c d::"'a::linordered_ab_group_add" proof - define wa where "wa = ua - la" define wb where "wb = ub - lb" define w where "w = wa + wb" define d where "d = x - la - lb" define da where "da = max 0 (min wa (d - wa))" define db where "db = d - da" from that have nonneg: "0 \ wa" "0 \ wb" "0 \ w" "0 \ d" "d \ w" by (auto simp add: wa_def wb_def w_def d_def add.commute le_diff_eq) have "0 \ db" by (auto simp: da_def nonneg db_def intro!: min.coboundedI2) have "x = (la + da) + (lb + db)" by (simp add: da_def db_def d_def) moreover have "x - la - ub \ da" using that unfolding da_def by (intro max.coboundedI2) (auto simp: wa_def d_def diff_le_eq diff_add_eq) then have "db \ wb" by (auto simp: db_def d_def wb_def algebra_simps) with \0 \ db\ that nonneg have "lb + db \ {lb..ub}" by (auto simp: wb_def algebra_simps) moreover have "da \ wa" by (auto simp: da_def nonneg) then have "la + da \ {la..ua}" by (auto simp: da_def wa_def algebra_simps) ultimately show ?thesis by force qed lemma ex_sum_in_interval: "\xa\la. xa \ ua \ (\xb\lb. xb \ ub \ x = xa + xb)" if a: "la \ ua" and b: "lb \ ub" and x: "la + lb \ x" "x \ ua + ub" for la b c d::"'a::linordered_ab_group_add" proof - from linear consider "ua - la \ ub - lb" | "ub - lb \ ua - la" by blast then show ?thesis proof cases case 1 from ex_sum_in_interval_lemma[OF that 1] show ?thesis by auto next case 2 from x have "lb + la \ x" "x \ ub + ua" by (simp_all add: ac_simps) from ex_sum_in_interval_lemma[OF b a this 2] show ?thesis by auto qed qed lemma Icc_plus_Icc: "{a .. b} + {c .. d} = {a + c .. b + d}" if "a \ b" "c \ d" for a b c d::"'a::linordered_ab_group_add" using ex_sum_in_interval[OF that] by (auto intro: add_mono simp: atLeastAtMost_iff Bex_def set_plus_def) lemma set_of_plus: fixes A :: "'a::linordered_ab_group_add interval" shows "set_of (A + B) = set_of A + set_of B" using Icc_plus_Icc[of "lower A" "upper A" "lower B" "upper B"] by (auto simp: set_of_eq) lemma plus_in_intervalE: fixes xy :: "'a :: linordered_ab_group_add" assumes "xy \\<^sub>i X + Y" obtains x y where "xy = x + y" "x \\<^sub>i X" "y \\<^sub>i Y" using assms unfolding set_of_plus set_plus_def by auto lemma set_of_uminus: "set_of (-X) = {- x | x. x \ set_of X}" for X :: "'a :: ordered_ab_group_add interval" by (auto simp: set_of_eq simp: le_minus_iff minus_le_iff intro!: exI[where x="-x" for x]) lemma uminus_in_intervalI: fixes x :: "'a :: ordered_ab_group_add" shows "x \\<^sub>i X \ -x \\<^sub>i -X" by (auto simp: set_of_uminus) lemma uminus_in_intervalD: fixes x :: "'a :: ordered_ab_group_add" shows "x \\<^sub>i - X \ - x \\<^sub>i X" by (auto simp: set_of_uminus) lemma minus_in_intervalI: fixes x y :: "'a :: ordered_ab_group_add" shows "x \\<^sub>i X \ y \\<^sub>i Y \ x - y \\<^sub>i X - Y" by (metis diff_conv_add_uminus minus_interval_def plus_in_intervalI uminus_in_intervalI) lemma set_of_minus: "set_of (X - Y) = {x - y | x y . x \ set_of X \ y \ set_of Y}" for X Y :: "'a :: linordered_ab_group_add interval" unfolding minus_interval_def set_of_plus set_of_uminus set_plus_def by force lemma times_in_intervalI: fixes x y::"'a::linordered_ring" assumes "x \\<^sub>i X" "y \\<^sub>i Y" shows "x * y \\<^sub>i X * Y" proof - define X1 where "X1 \ lower X" define X2 where "X2 \ upper X" define Y1 where "Y1 \ lower Y" define Y2 where "Y2 \ upper Y" from assms have assms: "X1 \ x" "x \ X2" "Y1 \ y" "y \ Y2" by (auto simp: X1_def X2_def Y1_def Y2_def set_of_eq) have "(X1 * Y1 \ x * y \ X1 * Y2 \ x * y \ X2 * Y1 \ x * y \ X2 * Y2 \ x * y) \ (X1 * Y1 \ x * y \ X1 * Y2 \ x * y \ X2 * Y1 \ x * y \ X2 * Y2 \ x * y)" proof (cases x "0::'a" rule: linorder_cases) case x0: less show ?thesis proof (cases "y < 0") case y0: True from y0 x0 assms have "x * y \ X1 * y" by (intro mult_right_mono_neg, auto) also from x0 y0 assms have "X1 * y \ X1 * Y1" by (intro mult_left_mono_neg, auto) finally have 1: "x * y \ X1 * Y1". show ?thesis proof(cases "X2 \ 0") case True with assms have "X2 * Y2 \ X2 * y" by (auto intro: mult_left_mono_neg) also from assms y0 have "... \ x * y" by (auto intro: mult_right_mono_neg) finally have "X2 * Y2 \ x * y". with 1 show ?thesis by auto next case False with assms have "X2 * Y1 \ X2 * y" by (auto intro: mult_left_mono) also from assms y0 have "... \ x * y" by (auto intro: mult_right_mono_neg) finally have "X2 * Y1 \ x * y". with 1 show ?thesis by auto qed next case False then have y0: "y \ 0" by auto from x0 y0 assms have "X1 * Y2 \ x * Y2" by (intro mult_right_mono, auto) also from y0 x0 assms have "... \ x * y" by (intro mult_left_mono_neg, auto) finally have 1: "X1 * Y2 \ x * y". show ?thesis proof(cases "X2 \ 0") case X2: True from assms y0 have "x * y \ X2 * y" by (intro mult_right_mono) also from assms X2 have "... \ X2 * Y1" by (auto intro: mult_left_mono_neg) finally have "x * y \ X2 * Y1". with 1 show ?thesis by auto next case X2: False from assms y0 have "x * y \ X2 * y" by (intro mult_right_mono) also from assms X2 have "... \ X2 * Y2" by (auto intro: mult_left_mono) finally have "x * y \ X2 * Y2". with 1 show ?thesis by auto qed qed next case [simp]: equal with assms show ?thesis by (cases "Y2 \ 0", auto intro:mult_sign_intros) next case x0: greater show ?thesis proof (cases "y < 0") case y0: True from x0 y0 assms have "X2 * Y1 \ X2 * y" by (intro mult_left_mono, auto) also from y0 x0 assms have "X2 * y \ x * y" by (intro mult_right_mono_neg, auto) finally have 1: "X2 * Y1 \ x * y". show ?thesis proof(cases "Y2 \ 0") case Y2: True from x0 assms have "x * y \ x * Y2" by (auto intro: mult_left_mono) also from assms Y2 have "... \ X1 * Y2" by (auto intro: mult_right_mono_neg) finally have "x * y \ X1 * Y2". with 1 show ?thesis by auto next case Y2: False from x0 assms have "x * y \ x * Y2" by (auto intro: mult_left_mono) also from assms Y2 have "... \ X2 * Y2" by (auto intro: mult_right_mono) finally have "x * y \ X2 * Y2". with 1 show ?thesis by auto qed next case y0: False from x0 y0 assms have "x * y \ X2 * y" by (intro mult_right_mono, auto) also from y0 x0 assms have "... \ X2 * Y2" by (intro mult_left_mono, auto) finally have 1: "x * y \ X2 * Y2". show ?thesis proof(cases "X1 \ 0") case True with assms have "X1 * Y2 \ X1 * y" by (auto intro: mult_left_mono_neg) also from assms y0 have "... \ x * y" by (auto intro: mult_right_mono) finally have "X1 * Y2 \ x * y". with 1 show ?thesis by auto next case False with assms have "X1 * Y1 \ X1 * y" by (auto intro: mult_left_mono) also from assms y0 have "... \ x * y" by (auto intro: mult_right_mono) finally have "X1 * Y1 \ x * y". with 1 show ?thesis by auto qed qed qed hence min:"min (X1 * Y1) (min (X1 * Y2) (min (X2 * Y1) (X2 * Y2))) \ x * y" and max:"x * y \ max (X1 * Y1) (max (X1 * Y2) (max (X2 * Y1) (X2 * Y2)))" by (auto simp:min_le_iff_disj le_max_iff_disj) show ?thesis using min max by (auto simp: Let_def X1_def X2_def Y1_def Y2_def set_of_eq lower_times upper_times) qed lemma times_in_intervalE: fixes xy :: "'a :: {linordered_semiring, real_normed_algebra, linear_continuum_topology}" \ \TODO: linear continuum topology is pretty strong\ assumes "xy \\<^sub>i X * Y" obtains x y where "xy = x * y" "x \\<^sub>i X" "y \\<^sub>i Y" proof - let ?mult = "\(x, y). x * y" let ?XY = "set_of X \ set_of Y" have cont: "continuous_on ?XY ?mult" by (auto intro!: tendsto_eq_intros simp: continuous_on_def split_beta') have conn: "connected (?mult ` ?XY)" by (rule connected_continuous_image[OF cont]) auto have "lower (X * Y) \ ?mult ` ?XY" "upper (X * Y) \ ?mult ` ?XY" by (auto simp: set_of_eq lower_times upper_times min_def max_def split: if_splits) from connectedD_interval[OF conn this, of xy] assms obtain x y where "xy = x * y" "x \\<^sub>i X" "y \\<^sub>i Y" by (auto simp: set_of_eq) then show ?thesis .. qed lemma set_of_times: "set_of (X * Y) = {x * y | x y. x \ set_of X \ y \ set_of Y}" for X Y::"'a :: {linordered_ring, real_normed_algebra, linear_continuum_topology} interval" by (auto intro!: times_in_intervalI elim!: times_in_intervalE) instance "interval" :: (linordered_idom) cancel_semigroup_add proof qed (auto simp: interval_eq_iff) lemma interval_mul_commute: "A * B = B * A" for A B:: "'a::linordered_idom interval" by (simp add: interval_eq_iff lower_times upper_times ac_simps) lemma interval_times_zero_right[simp]: "A * 0 = 0" for A :: "'a::linordered_ring interval" by (simp add: interval_eq_iff lower_times upper_times ac_simps) lemma interval_times_zero_left[simp]: "0 * A = 0" for A :: "'a::linordered_ring interval" by (simp add: interval_eq_iff lower_times upper_times ac_simps) instantiation "interval" :: ("{preorder,one}") one begin lift_definition one_interval::"'a interval" is "(1, 1)" by auto lemma lower_one[simp]: "lower 1 = 1" by transfer auto lemma upper_one[simp]: "upper 1 = 1" by transfer auto instance proof qed end instance interval :: ("{one, preorder, linordered_semiring}") power proof qed lemma set_of_one[simp]: "set_of (1::'a::{one, order} interval) = {1}" by (auto simp: set_of_eq) instance "interval" :: ("{linordered_idom,linordered_ring, real_normed_algebra, linear_continuum_topology}") monoid_mult apply standard unfolding interval_eq_set_of_iff set_of_times subgoal by (auto simp: interval_eq_set_of_iff set_of_times; metis mult.assoc) by auto lemma one_times_ivl_left[simp]: "1 * A = A" for A :: "'a::linordered_idom interval" by (simp add: interval_eq_iff lower_times upper_times ac_simps min_def max_def) lemma one_times_ivl_right[simp]: "A * 1 = A" for A :: "'a::linordered_idom interval" by (metis interval_mul_commute one_times_ivl_left) lemma set_of_power_mono: "a^n \ set_of (A^n)" if "a \ set_of A" for a :: "'a::linordered_idom" using that by (induction n) (auto intro!: times_in_intervalI) lemma set_of_add_cong: "set_of (A + B) = set_of (A' + B')" if "set_of A = set_of A'" "set_of B = set_of B'" for A :: "'a::linordered_ab_group_add interval" unfolding set_of_plus that .. lemma set_of_add_inc_left: "set_of (A + B) \ set_of (A' + B)" if "set_of A \ set_of A'" for A :: "'a::linordered_ab_group_add interval" unfolding set_of_plus using that by (auto simp: set_plus_def) lemma set_of_add_inc_right: "set_of (A + B) \ set_of (A + B')" if "set_of B \ set_of B'" for A :: "'a::linordered_ab_group_add interval" using set_of_add_inc_left[OF that] by (simp add: add.commute) lemma set_of_add_inc: "set_of (A + B) \ set_of (A' + B')" if "set_of A \ set_of A'" "set_of B \ set_of B'" for A :: "'a::linordered_ab_group_add interval" using set_of_add_inc_left[OF that(1)] set_of_add_inc_right[OF that(2)] by auto lemma set_of_neg_inc: "set_of (-A) \ set_of (-A')" if "set_of A \ set_of A'" for A :: "'a::ordered_ab_group_add interval" using that unfolding set_of_uminus by auto lemma set_of_sub_inc_left: "set_of (A - B) \ set_of (A' - B)" if "set_of A \ set_of A'" for A :: "'a::linordered_ab_group_add interval" using that unfolding set_of_minus by auto lemma set_of_sub_inc_right: "set_of (A - B) \ set_of (A - B')" if "set_of B \ set_of B'" for A :: "'a::linordered_ab_group_add interval" using that unfolding set_of_minus by auto lemma set_of_sub_inc: "set_of (A - B) \ set_of (A' - B')" if "set_of A \ set_of A'" "set_of B \ set_of B'" for A :: "'a::linordered_idom interval" using set_of_sub_inc_left[OF that(1)] set_of_sub_inc_right[OF that(2)] by auto lemma set_of_mul_inc_right: "set_of (A * B) \ set_of (A * B')" if "set_of B \ set_of B'" for A :: "'a::linordered_ring interval" using that apply transfer apply (clarsimp simp add: Let_def) apply (intro conjI) apply (metis linear min.coboundedI1 min.coboundedI2 mult_left_mono mult_left_mono_neg order_trans) apply (metis linear min.coboundedI1 min.coboundedI2 mult_left_mono mult_left_mono_neg order_trans) apply (metis linear min.coboundedI1 min.coboundedI2 mult_left_mono mult_left_mono_neg order_trans) apply (metis linear min.coboundedI1 min.coboundedI2 mult_left_mono mult_left_mono_neg order_trans) apply (metis linear max.coboundedI1 max.coboundedI2 mult_left_mono mult_left_mono_neg order_trans) apply (metis linear max.coboundedI1 max.coboundedI2 mult_left_mono mult_left_mono_neg order_trans) apply (metis linear max.coboundedI1 max.coboundedI2 mult_left_mono mult_left_mono_neg order_trans) apply (metis linear max.coboundedI1 max.coboundedI2 mult_left_mono mult_left_mono_neg order_trans) done lemma set_of_distrib_left: "set_of (B * (A1 + A2)) \ set_of (B * A1 + B * A2)" for A1 :: "'a::linordered_ring interval" apply transfer apply (clarsimp simp: Let_def distrib_left distrib_right) apply (intro conjI) apply (metis add_mono min.cobounded1 min.left_commute) apply (metis add_mono min.cobounded1 min.left_commute) apply (metis add_mono min.cobounded1 min.left_commute) apply (metis add_mono min.assoc min.cobounded2) apply (meson add_mono order.trans max.cobounded1 max.cobounded2) apply (meson add_mono order.trans max.cobounded1 max.cobounded2) apply (meson add_mono order.trans max.cobounded1 max.cobounded2) apply (meson add_mono order.trans max.cobounded1 max.cobounded2) done lemma set_of_distrib_right: "set_of ((A1 + A2) * B) \ set_of (A1 * B + A2 * B)" for A1 A2 B :: "'a::{linordered_ring, real_normed_algebra, linear_continuum_topology} interval" unfolding set_of_times set_of_plus set_plus_def apply clarsimp subgoal for b a1 a2 apply (rule exI[where x="a1 * b"]) apply (rule conjI) subgoal by force subgoal apply (rule exI[where x="a2 * b"]) apply (rule conjI) subgoal by force subgoal by (simp add: algebra_simps) done done done lemma set_of_mul_inc_left: "set_of (A * B) \ set_of (A' * B)" if "set_of A \ set_of A'" for A :: "'a::{linordered_ring, real_normed_algebra, linear_continuum_topology} interval" using that unfolding set_of_times by auto lemma set_of_mul_inc: "set_of (A * B) \ set_of (A' * B')" if "set_of A \ set_of A'" "set_of B \ set_of B'" for A :: "'a::{linordered_ring, real_normed_algebra, linear_continuum_topology} interval" using that unfolding set_of_times by auto lemma set_of_pow_inc: "set_of (A^n) \ set_of (A'^n)" if "set_of A \ set_of A'" for A :: "'a::{linordered_idom, real_normed_algebra, linear_continuum_topology} interval" using that by (induction n, simp_all add: set_of_mul_inc) lemma set_of_distrib_right_left: "set_of ((A1 + A2) * (B1 + B2)) \ set_of (A1 * B1 + A1 * B2 + A2 * B1 + A2 * B2)" for A1 :: "'a::{linordered_idom, real_normed_algebra, linear_continuum_topology} interval" proof- have "set_of ((A1 + A2) * (B1 + B2)) \ set_of (A1 * (B1 + B2) + A2 * (B1 + B2))" by (rule set_of_distrib_right) also have "... \ set_of ((A1 * B1 + A1 * B2) + A2 * (B1 + B2))" by (rule set_of_add_inc_left[OF set_of_distrib_left]) also have "... \ set_of ((A1 * B1 + A1 * B2) + (A2 * B1 + A2 * B2))" by (rule set_of_add_inc_right[OF set_of_distrib_left]) finally show ?thesis by (simp add: add.assoc) qed lemma mult_bounds_enclose_zero1: "min (la * lb) (min (la * ub) (min (lb * ua) (ua * ub))) \ 0" "0 \ max (la * lb) (max (la * ub) (max (lb * ua) (ua * ub)))" if "la \ 0" "0 \ ua" for la lb ua ub:: "'a::linordered_idom" subgoal by (metis (no_types, opaque_lifting) that eq_iff min_le_iff_disj mult_zero_left mult_zero_right zero_le_mult_iff) subgoal by (metis that le_max_iff_disj mult_zero_right order_refl zero_le_mult_iff) done lemma mult_bounds_enclose_zero2: "min (la * lb) (min (la * ub) (min (lb * ua) (ua * ub))) \ 0" "0 \ max (la * lb) (max (la * ub) (max (lb * ua) (ua * ub)))" if "lb \ 0" "0 \ ub" for la lb ua ub:: "'a::linordered_idom" using mult_bounds_enclose_zero1[OF that, of la ua] by (simp_all add: ac_simps) lemma set_of_mul_contains_zero: "0 \ set_of (A * B)" if "0 \ set_of A \ 0 \ set_of B" for A :: "'a::linordered_idom interval" using that by (auto simp: set_of_eq lower_times upper_times algebra_simps mult_le_0_iff mult_bounds_enclose_zero1 mult_bounds_enclose_zero2) instance "interval" :: (linordered_semiring) mult_zero apply standard subgoal by transfer auto subgoal by transfer auto done lift_definition min_interval::"'a::linorder interval \ 'a interval \ 'a interval" is "\(l1, u1). \(l2, u2). (min l1 l2, min u1 u2)" by (auto simp: min_def) lemma lower_min_interval[simp]: "lower (min_interval x y) = min (lower x) (lower y)" by transfer auto lemma upper_min_interval[simp]: "upper (min_interval x y) = min (upper x) (upper y)" by transfer auto lemma min_intervalI: "a \\<^sub>i A \ b \\<^sub>i B \ min a b \\<^sub>i min_interval A B" by (auto simp: set_of_eq min_def) lift_definition max_interval::"'a::linorder interval \ 'a interval \ 'a interval" is "\(l1, u1). \(l2, u2). (max l1 l2, max u1 u2)" by (auto simp: max_def) lemma lower_max_interval[simp]: "lower (max_interval x y) = max (lower x) (lower y)" by transfer auto lemma upper_max_interval[simp]: "upper (max_interval x y) = max (upper x) (upper y)" by transfer auto lemma max_intervalI: "a \\<^sub>i A \ b \\<^sub>i B \ max a b \\<^sub>i max_interval A B" by (auto simp: set_of_eq max_def) lift_definition abs_interval::"'a::linordered_idom interval \ 'a interval" is "(\(l,u). (if l < 0 \ 0 < u then 0 else min \l\ \u\, max \l\ \u\))" by auto lemma lower_abs_interval[simp]: "lower (abs_interval x) = (if lower x < 0 \ 0 < upper x then 0 else min \lower x\ \upper x\)" by transfer auto lemma upper_abs_interval[simp]: "upper (abs_interval x) = max \lower x\ \upper x\" by transfer auto lemma in_abs_intervalI1: "lx < 0 \ 0 < ux \ 0 \ xa \ xa \ max (- lx) (ux) \ xa \ abs ` {lx..ux}" for xa::"'a::linordered_idom" by (metis abs_minus_cancel abs_of_nonneg atLeastAtMost_iff image_eqI le_less le_max_iff_disj le_minus_iff neg_le_0_iff_le order_trans) lemma in_abs_intervalI2: "min (\lx\) \ux\ \ xa \ xa \ max \lx\ \ux\ \ lx \ ux \ 0 \ lx \ ux \ 0 \ xa \ abs ` {lx..ux}" for xa::"'a::linordered_idom" by (force intro: image_eqI[where x="-xa"] image_eqI[where x="xa"]) lemma set_of_abs_interval: "set_of (abs_interval x) = abs ` set_of x" by (auto simp: set_of_eq not_less intro: in_abs_intervalI1 in_abs_intervalI2 cong del: image_cong_simp) fun split_domain :: "('a::preorder interval \ 'a interval list) \ 'a interval list \ 'a interval list list" where "split_domain split [] = [[]]" | "split_domain split (I#Is) = ( let S = split I; D = split_domain split Is in concat (map (\d. map (\s. s # d) S) D) )" context notes [[typedef_overloaded]] begin lift_definition(code_dt) split_interval::"'a::linorder interval \ 'a \ ('a interval \ 'a interval)" is "\(l, u) x. ((min l x, max l x), (min u x, max u x))" by (auto simp: min_def) end lemma split_domain_nonempty: assumes "\I. split I \ []" shows "split_domain split I \ []" using last_in_set assms by (induction I, auto) lemma lower_split_interval1: "lower (fst (split_interval X m)) = min (lower X) m" and lower_split_interval2: "lower (snd (split_interval X m)) = min (upper X) m" and upper_split_interval1: "upper (fst (split_interval X m)) = max (lower X) m" and upper_split_interval2: "upper (snd (split_interval X m)) = max (upper X) m" subgoal by transfer auto subgoal by transfer (auto simp: min.commute) - subgoal by transfer (auto simp: ) - subgoal by transfer (auto simp: ) + subgoal by transfer auto + subgoal by transfer auto done lemma split_intervalD: "split_interval X x = (A, B) \ set_of X \ set_of A \ set_of B" unfolding set_of_eq by transfer (auto simp: min_def max_def split: if_splits) instantiation interval :: ("{topological_space, preorder}") topological_space begin definition open_interval_def[code del]: "open (X::'a interval set) = (\x\X. \A B. open A \ open B \ lower x \ A \ upper x \ B \ Interval ` (A \ B) \ X)" instance proof show "open (UNIV :: ('a interval) set)" unfolding open_interval_def by auto next fix S T :: "('a interval) set" assume "open S" "open T" show "open (S \ T)" unfolding open_interval_def proof (safe) fix x assume "x \ S" "x \ T" from \x \ S\ \open S\ obtain Sl Su where S: "open Sl" "open Su" "lower x \ Sl" "upper x \ Su" "Interval ` (Sl \ Su) \ S" by (auto simp: open_interval_def) from \x \ T\ \open T\ obtain Tl Tu where T: "open Tl" "open Tu" "lower x \ Tl" "upper x \ Tu" "Interval ` (Tl \ Tu) \ T" by (auto simp: open_interval_def) let ?L = "Sl \ Tl" and ?U = "Su \ Tu" have "open ?L \ open ?U \ lower x \ ?L \ upper x \ ?U \ Interval ` (?L \ ?U) \ S \ T" using S T by (auto simp add: open_Int) then show "\A B. open A \ open B \ lower x \ A \ upper x \ B \ Interval ` (A \ B) \ S \ T" by fast qed qed (unfold open_interval_def, fast) end subsection \Quickcheck\ lift_definition Ivl::"'a \ 'a::preorder \ 'a interval" is "\a b. (min a b, b)" by (auto simp: min_def) instantiation interval :: ("{exhaustive,preorder}") exhaustive begin definition exhaustive_interval::"('a interval \ (bool \ term list) option) \ natural \ (bool \ term list) option" where "exhaustive_interval f d = Quickcheck_Exhaustive.exhaustive (\x. Quickcheck_Exhaustive.exhaustive (\y. f (Ivl x y)) d) d" instance .. end context includes term_syntax begin definition [code_unfold]: "valtermify_interval x y = Code_Evaluation.valtermify (Ivl::'a::{preorder,typerep}\_) {\} x {\} y" end instantiation interval :: ("{full_exhaustive,preorder,typerep}") full_exhaustive begin definition full_exhaustive_interval:: "('a interval \ (unit \ term) \ (bool \ term list) option) \ natural \ (bool \ term list) option" where "full_exhaustive_interval f d = Quickcheck_Exhaustive.full_exhaustive (\x. Quickcheck_Exhaustive.full_exhaustive (\y. f (valtermify_interval x y)) d) d" instance .. end instantiation interval :: ("{random,preorder,typerep}") random begin definition random_interval :: "natural \ natural \ natural \ ('a interval \ (unit \ term)) \ natural \ natural" where "random_interval i = scomp (Quickcheck_Random.random i) (\man. scomp (Quickcheck_Random.random i) (\exp. Pair (valtermify_interval man exp)))" instance .. end lifting_update interval.lifting lifting_forget interval.lifting end diff --git a/src/HOL/Library/Log_Nat.thy b/src/HOL/Library/Log_Nat.thy --- a/src/HOL/Library/Log_Nat.thy +++ b/src/HOL/Library/Log_Nat.thy @@ -1,288 +1,288 @@ (* Title: HOL/Library/Log_Nat.thy Author: Johannes Hölzl, Fabian Immler Copyright 2012 TU München *) section \Logarithm of Natural Numbers\ theory Log_Nat imports Complex_Main begin subsection \Preliminaries\ lemma divide_nat_diff_div_nat_less_one: "real x / real b - real (x div b) < 1" for x b :: nat proof (cases "b = 0") case True then show ?thesis by simp next case False then have "real (x div b) + real (x mod b) / real b - real (x div b) < 1" by (simp add: field_simps) then show ?thesis by (simp add: real_of_nat_div_aux [symmetric]) qed subsection \Floorlog\ definition floorlog :: "nat \ nat \ nat" where "floorlog b a = (if a > 0 \ b > 1 then nat \log b a\ + 1 else 0)" lemma floorlog_mono: "x \ y \ floorlog b x \ floorlog b y" by (auto simp: floorlog_def floor_mono nat_mono) lemma floorlog_bounds: "b ^ (floorlog b x - 1) \ x \ x < b ^ (floorlog b x)" if "x > 0" "b > 1" proof show "b ^ (floorlog b x - 1) \ x" proof - have "b ^ nat \log b x\ = b powr \log b x\" using powr_realpow[symmetric, of b "nat \log b x\"] \x > 0\ \b > 1\ by simp also have "\ \ b powr log b x" using \b > 1\ by simp also have "\ = real_of_int x" using \0 < x\ \b > 1\ by simp finally have "b ^ nat \log b x\ \ real_of_int x" by simp then show ?thesis using \0 < x\ \b > 1\ of_nat_le_iff by (fastforce simp add: floorlog_def) qed show "x < b ^ (floorlog b x)" proof - have "x \ b powr (log b x)" using \x > 0\ \b > 1\ by simp also have "\ < b powr (\log b x\ + 1)" using that by (intro powr_less_mono) auto also have "\ = b ^ nat (\log b (real_of_int x)\ + 1)" using that by (simp flip: powr_realpow) finally have "x < b ^ nat (\log b (int x)\ + 1)" by (rule of_nat_less_imp_less) then show ?thesis using \x > 0\ \b > 1\ by (simp add: floorlog_def nat_add_distrib) qed qed lemma floorlog_power [simp]: "floorlog b (a * b ^ c) = floorlog b a + c" if "a > 0" "b > 1" proof - have "\log b a + real c\ = \log b a\ + c" by arith then show ?thesis using that by (auto simp: floorlog_def log_mult powr_realpow[symmetric] nat_add_distrib) qed lemma floor_log_add_eqI: "\log b (a + r)\ = \log b a\" if "b > 1" "a \ 1" "0 \ r" "r < 1" for a b :: nat and r :: real proof (rule floor_eq2) have "log b a \ log b (a + r)" using that by force then show "\log b a\ \ log b (a + r)" by arith next define l::int where "l = int b ^ (nat \log b a\ + 1)" have l_def_real: "l = b powr (\log b a\ + 1)" using that by (simp add: l_def powr_add powr_real_of_int) have "a < l" proof - have "a = b powr (log b a)" using that by simp also have "\ < b powr floor ((log b a) + 1)" using that(1) by auto also have "\ = l" using that by (simp add: l_def powr_real_of_int powr_add) finally show ?thesis by simp qed then have "a + r < l" using that by simp then have "log b (a + r) < log b l" using that by simp also have "\ = real_of_int \log b a\ + 1" using that by (simp add: l_def_real) finally show "log b (a + r) < real_of_int \log b a\ + 1" . qed lemma floor_log_div: "\log b x\ = \log b (x div b)\ + 1" if "b > 1" "x > 0" "x div b > 0" for b x :: nat proof- have "\log b x\ = \log b (x / b * b)\" using that by simp also have "\ = \log b (x / b) + log b b\" using that by (subst log_mult) auto also have "\ = \log b (x / b)\ + 1" using that by simp also have "\log b (x / b)\ = \log b (x div b + (x / b - x div b))\" by simp also have "\ = \log b (x div b)\" using that real_of_nat_div4 divide_nat_diff_div_nat_less_one by (intro floor_log_add_eqI) auto finally show ?thesis . qed lemma compute_floorlog [code]: "floorlog b x = (if x > 0 \ b > 1 then floorlog b (x div b) + 1 else 0)" by (auto simp: floorlog_def floor_log_div[of b x] div_eq_0_iff nat_add_distrib intro!: floor_eq2) lemma floor_log_eq_if: "\log b x\ = \log b y\" if "x div b = y div b" "b > 1" "x > 0" "x div b \ 1" for b x y :: nat proof - have "y > 0" using that by (auto intro: ccontr) thus ?thesis using that by (simp add: floor_log_div) qed lemma floorlog_eq_if: "floorlog b x = floorlog b y" if "x div b = y div b" "b > 1" "x > 0" "x div b \ 1" for b x y :: nat proof - have "y > 0" using that by (auto intro: ccontr) then show ?thesis using that by (auto simp add: floorlog_def eq_nat_nat_iff intro: floor_log_eq_if) qed lemma floorlog_leD: "floorlog b x \ w \ b > 1 \ x < b ^ w" by (metis floorlog_bounds leD linorder_neqE_nat order.strict_trans power_strict_increasing_iff zero_less_one zero_less_power) lemma floorlog_leI: "x < b ^ w \ 0 \ w \ b > 1 \ floorlog b x \ w" by (drule less_imp_of_nat_less[where 'a=real]) (auto simp: floorlog_def Suc_le_eq nat_less_iff floor_less_iff log_of_power_less) lemma floorlog_eq_zero_iff: "floorlog b x = 0 \ b \ 1 \ x \ 0" by (auto simp: floorlog_def) lemma floorlog_le_iff: "floorlog b x \ w \ b \ 1 \ b > 1 \ 0 \ w \ x < b ^ w" using floorlog_leD[of b x w] floorlog_leI[of x b w] by (auto simp: floorlog_eq_zero_iff[THEN iffD2]) lemma floorlog_ge_SucI: "Suc w \ floorlog b x" if "b ^ w \ x" "b > 1" using that le_log_of_power[of b w x] power_not_zero by (force simp: floorlog_def Suc_le_eq powr_realpow not_less Suc_nat_eq_nat_zadd1 zless_nat_eq_int_zless int_add_floor less_floor_iff simp del: floor_add2) lemma floorlog_geI: "w \ floorlog b x" if "b ^ (w - 1) \ x" "b > 1" using floorlog_ge_SucI[of b "w - 1" x] that by auto lemma floorlog_geD: "b ^ (w - 1) \ x" if "w \ floorlog b x" "w > 0" proof - have "b > 1" "0 < x" using that by (auto simp: floorlog_def split: if_splits) have "b ^ (w - 1) \ x" if "b ^ w \ x" proof - have "b ^ (w - 1) \ b ^ w" using \b > 1\ by (auto intro!: power_increasing) also note that finally show ?thesis . qed moreover have "b ^ nat \log (real b) (real x)\ \ x" (is "?l \ _") proof - have "0 \ log (real b) (real x)" using \b > 1\ \0 < x\ - by (auto simp: ) + by auto then have "?l \ b powr log (real b) (real x)" using \b > 1\ by (auto simp flip: powr_realpow intro!: powr_mono of_nat_floor) also have "\ = x" using \b > 1\ \0 < x\ by auto finally show ?thesis unfolding of_nat_le_iff . qed ultimately show ?thesis using that by (auto simp: floorlog_def le_nat_iff le_floor_iff le_log_iff powr_realpow split: if_splits elim!: le_SucE) qed subsection \Bitlen\ definition bitlen :: "int \ int" where "bitlen a = floorlog 2 (nat a)" lemma bitlen_alt_def: "bitlen a = (if a > 0 then \log 2 a\ + 1 else 0)" by (simp add: bitlen_def floorlog_def) lemma bitlen_zero [simp]: "bitlen 0 = 0" by (auto simp: bitlen_def floorlog_def) lemma bitlen_nonneg: "0 \ bitlen x" by (simp add: bitlen_def) lemma bitlen_bounds: "2 ^ nat (bitlen x - 1) \ x \ x < 2 ^ nat (bitlen x)" if "x > 0" proof - from that have "bitlen x \ 1" by (auto simp: bitlen_alt_def) with that floorlog_bounds[of "nat x" 2] show ?thesis by (auto simp add: bitlen_def le_nat_iff nat_less_iff nat_diff_distrib) qed lemma bitlen_pow2 [simp]: "bitlen (b * 2 ^ c) = bitlen b + c" if "b > 0" using that by (simp add: bitlen_def nat_mult_distrib nat_power_eq) lemma compute_bitlen [code]: "bitlen x = (if x > 0 then bitlen (x div 2) + 1 else 0)" by (simp add: bitlen_def nat_div_distrib compute_floorlog) lemma bitlen_eq_zero_iff: "bitlen x = 0 \ x \ 0" by (auto simp add: bitlen_alt_def) (metis compute_bitlen add.commute bitlen_alt_def bitlen_nonneg less_add_same_cancel2 not_less zero_less_one) lemma bitlen_div: "1 \ real_of_int m / 2^nat (bitlen m - 1)" and "real_of_int m / 2^nat (bitlen m - 1) < 2" if "0 < m" proof - let ?B = "2^nat (bitlen m - 1)" have "?B \ m" using bitlen_bounds[OF \0 ] .. then have "1 * ?B \ real_of_int m" unfolding of_int_le_iff[symmetric] by auto then show "1 \ real_of_int m / ?B" by auto from that have "0 \ bitlen m - 1" by (auto simp: bitlen_alt_def) have "m < 2^nat(bitlen m)" using bitlen_bounds[OF that] .. also from that have "\ = 2^nat(bitlen m - 1 + 1)" by (auto simp: bitlen_def) also have "\ = ?B * 2" unfolding nat_add_distrib[OF \0 \ bitlen m - 1\ zero_le_one] by auto finally have "real_of_int m < 2 * ?B" by (metis (full_types) mult.commute power.simps(2) of_int_less_numeral_power_cancel_iff) then have "real_of_int m / ?B < 2 * ?B / ?B" by (rule divide_strict_right_mono) auto then show "real_of_int m / ?B < 2" by auto qed lemma bitlen_le_iff_floorlog: "bitlen x \ w \ w \ 0 \ floorlog 2 (nat x) \ nat w" by (auto simp: bitlen_def) lemma bitlen_le_iff_power: "bitlen x \ w \ w \ 0 \ x < 2 ^ nat w" by (auto simp: bitlen_le_iff_floorlog floorlog_le_iff) lemma less_power_nat_iff_bitlen: "x < 2 ^ w \ bitlen (int x) \ w" using bitlen_le_iff_power[of x w] by auto lemma bitlen_ge_iff_power: "w \ bitlen x \ w \ 0 \ 2 ^ (nat w - 1) \ x" unfolding bitlen_def by (auto simp flip: nat_le_iff intro: floorlog_geI dest: floorlog_geD) lemma bitlen_twopow_add_eq: "bitlen (2 ^ w + b) = w + 1" if "0 \ b" "b < 2 ^ w" by (auto simp: that nat_add_distrib bitlen_le_iff_power bitlen_ge_iff_power intro!: antisym) end diff --git a/src/HOL/Library/Poly_Mapping.thy b/src/HOL/Library/Poly_Mapping.thy --- a/src/HOL/Library/Poly_Mapping.thy +++ b/src/HOL/Library/Poly_Mapping.thy @@ -1,1878 +1,1878 @@ (* Author: Andreas Lochbihler, ETH Zurich Author: Florian Haftmann, TU Muenchen with some material ported from HOL Light by LCP *) section \Polynomial mapping: combination of almost everywhere zero functions with an algebraic view\ theory Poly_Mapping imports Groups_Big_Fun Fun_Lexorder More_List begin subsection \Preliminary: auxiliary operations for \emph{almost everywhere zero}\ text \ A central notion for polynomials are functions being \emph{almost everywhere zero}. For these we provide some auxiliary definitions and lemmas. \ lemma finite_mult_not_eq_zero_leftI: fixes f :: "'b \ 'a :: mult_zero" assumes "finite {a. f a \ 0}" shows "finite {a. g a * f a \ 0}" proof - have "{a. g a * f a \ 0} \ {a. f a \ 0}" by auto then show ?thesis using assms by (rule finite_subset) qed lemma finite_mult_not_eq_zero_rightI: fixes f :: "'b \ 'a :: mult_zero" assumes "finite {a. f a \ 0}" shows "finite {a. f a * g a \ 0}" proof - have "{a. f a * g a \ 0} \ {a. f a \ 0}" by auto then show ?thesis using assms by (rule finite_subset) qed lemma finite_mult_not_eq_zero_prodI: fixes f g :: "'a \ 'b::semiring_0" assumes "finite {a. f a \ 0}" (is "finite ?F") assumes "finite {b. g b \ 0}" (is "finite ?G") shows "finite {(a, b). f a * g b \ 0}" proof - from assms have "finite (?F \ ?G)" by blast then have "finite {(a, b). f a \ 0 \ g b \ 0}" by simp then show ?thesis by (rule rev_finite_subset) auto qed lemma finite_not_eq_zero_sumI: fixes f g :: "'a::monoid_add \ 'b::semiring_0" assumes "finite {a. f a \ 0}" (is "finite ?F") assumes "finite {b. g b \ 0}" (is "finite ?G") shows "finite {a + b | a b. f a \ 0 \ g b \ 0}" (is "finite ?FG") proof - from assms have "finite (?F \ ?G)" by (simp add: finite_cartesian_product_iff) then have "finite (case_prod plus ` (?F \ ?G))" by (rule finite_imageI) also have "case_prod plus ` (?F \ ?G) = ?FG" by auto finally show ?thesis by simp qed lemma finite_mult_not_eq_zero_sumI: fixes f g :: "'a::monoid_add \ 'b::semiring_0" assumes "finite {a. f a \ 0}" assumes "finite {b. g b \ 0}" shows "finite {a + b | a b. f a * g b \ 0}" proof - from assms have "finite {a + b | a b. f a \ 0 \ g b \ 0}" by (rule finite_not_eq_zero_sumI) then show ?thesis by (rule rev_finite_subset) (auto dest: mult_not_zero) qed lemma finite_Sum_any_not_eq_zero_weakenI: assumes "finite {a. \b. f a b \ 0}" shows "finite {a. Sum_any (f a) \ 0}" proof - have "{a. Sum_any (f a) \ 0} \ {a. \b. f a b \ 0}" by (auto elim: Sum_any.not_neutral_obtains_not_neutral) then show ?thesis using assms by (rule finite_subset) qed context zero begin definition "when" :: "'a \ bool \ 'a" (infixl "when" 20) where "(a when P) = (if P then a else 0)" text \ Case distinctions always complicate matters, particularly when nested. The @{const "when"} operation allows to minimise these if @{term 0} is the false-case value and makes proof obligations much more readable. \ lemma "when" [simp]: "P \ (a when P) = a" "\ P \ (a when P) = 0" by (simp_all add: when_def) lemma when_simps [simp]: "(a when True) = a" "(a when False) = 0" by simp_all lemma when_cong: assumes "P \ Q" and "Q \ a = b" shows "(a when P) = (b when Q)" using assms by (simp add: when_def) lemma zero_when [simp]: "(0 when P) = 0" by (simp add: when_def) lemma when_when: "(a when P when Q) = (a when P \ Q)" by (cases Q) simp_all lemma when_commute: "(a when Q when P) = (a when P when Q)" by (simp add: when_when conj_commute) lemma when_neq_zero [simp]: "(a when P) \ 0 \ P \ a \ 0" by (cases P) simp_all end context monoid_add begin lemma when_add_distrib: "(a + b when P) = (a when P) + (b when P)" by (simp add: when_def) end context semiring_1 begin lemma zero_power_eq: "0 ^ n = (1 when n = 0)" by (simp add: power_0_left) end context comm_monoid_add begin lemma Sum_any_when_equal [simp]: "(\a. (f a when a = b)) = f b" by (simp add: when_def) lemma Sum_any_when_equal' [simp]: "(\a. (f a when b = a)) = f b" by (simp add: when_def) lemma Sum_any_when_independent: "(\a. g a when P) = ((\a. g a) when P)" by (cases P) simp_all lemma Sum_any_when_dependent_prod_right: "(\(a, b). g a when b = h a) = (\a. g a)" proof - have "inj_on (\a. (a, h a)) {a. g a \ 0}" by (rule inj_onI) auto then show ?thesis unfolding Sum_any.expand_set by (rule sum.reindex_cong) auto qed lemma Sum_any_when_dependent_prod_left: "(\(a, b). g b when a = h b) = (\b. g b)" proof - have "(\(a, b). g b when a = h b) = (\(b, a). g b when a = h b)" by (rule Sum_any.reindex_cong [of prod.swap]) (simp_all add: fun_eq_iff) then show ?thesis by (simp add: Sum_any_when_dependent_prod_right) qed end context cancel_comm_monoid_add begin lemma when_diff_distrib: "(a - b when P) = (a when P) - (b when P)" by (simp add: when_def) end context group_add begin lemma when_uminus_distrib: "(- a when P) = - (a when P)" by (simp add: when_def) end context mult_zero begin lemma mult_when: "a * (b when P) = (a * b when P)" by (cases P) simp_all lemma when_mult: "(a when P) * b = (a * b when P)" by (cases P) simp_all end subsection \Type definition\ text \ The following type is of central importance: \ typedef (overloaded) ('a, 'b) poly_mapping ("(_ \\<^sub>0 /_)" [1, 0] 0) = "{f :: 'a \ 'b::zero. finite {x. f x \ 0}}" morphisms lookup Abs_poly_mapping proof - have "(\_::'a. (0 :: 'b)) \ ?poly_mapping" by simp then show ?thesis by (blast intro!: exI) qed declare lookup_inverse [simp] declare lookup_inject [simp] lemma lookup_Abs_poly_mapping [simp]: "finite {x. f x \ 0} \ lookup (Abs_poly_mapping f) = f" using Abs_poly_mapping_inverse [of f] by simp lemma finite_lookup [simp]: "finite {k. lookup f k \ 0}" using poly_mapping.lookup [of f] by simp lemma finite_lookup_nat [simp]: fixes f :: "'a \\<^sub>0 nat" shows "finite {k. 0 < lookup f k}" using poly_mapping.lookup [of f] by simp lemma poly_mapping_eqI: assumes "\k. lookup f k = lookup g k" shows "f = g" using assms unfolding poly_mapping.lookup_inject [symmetric] by blast lemma poly_mapping_eq_iff: "a = b \ lookup a = lookup b" by auto text \ We model the universe of functions being \emph{almost everywhere zero} by means of a separate type @{typ "('a, 'b) poly_mapping"}. For convenience we provide a suggestive infix syntax which is a variant of the usual function space syntax. Conversion between both types happens through the morphisms \begin{quote} @{term_type lookup} \end{quote} \begin{quote} @{term_type Abs_poly_mapping} \end{quote} satisfying \begin{quote} @{thm lookup_inverse} \end{quote} \begin{quote} @{thm lookup_Abs_poly_mapping} \end{quote} Luckily, we have rarely to deal with those low-level morphisms explicitly but rely on Isabelle's \emph{lifting} package with its method \transfer\ and its specification tool \lift_definition\. \ setup_lifting type_definition_poly_mapping code_datatype Abs_poly_mapping\\FIXME? workaround for preventing \code_abstype\ setup\ text \ @{typ "'a \\<^sub>0 'b"} serves distinctive purposes: \begin{enumerate} \item A clever nesting as @{typ "(nat \\<^sub>0 nat) \\<^sub>0 'a"} later in theory \MPoly\ gives a suitable representation type for polynomials \emph{almost for free}: Interpreting @{typ "nat \\<^sub>0 nat"} as a mapping from variable identifiers to exponents yields monomials, and the whole type maps monomials to coefficients. Lets call this the \emph{ultimate interpretation}. \item A further more specialised type isomorphic to @{typ "nat \\<^sub>0 'a"} is apt to direct implementation using code generation \cite{Haftmann-Nipkow:2010:code}. \end{enumerate} Note that despite the names \emph{mapping} and \emph{lookup} suggest something implementation-near, it is best to keep @{typ "'a \\<^sub>0 'b"} as an abstract \emph{algebraic} type providing operations like \emph{addition}, \emph{multiplication} without any notion of key-order, data structures etc. This implementations-specific notions are easily introduced later for particular implementations but do not provide any gain for specifying logical properties of polynomials. \ subsection \Additive structure\ text \ The additive structure covers the usual operations \0\, \+\ and (unary and binary) \-\. Recalling the ultimate interpretation, it is obvious that these have just lift the corresponding operations on values to mappings. Isabelle has a rich hierarchy of algebraic type classes, and in such situations of pointwise lifting a typical pattern is to have instantiations for a considerable number of type classes. The operations themselves are specified using \lift_definition\, where the proofs of the \emph{almost everywhere zero} property can be significantly involved. The @{const lookup} operation is supposed to be usable explicitly (unless in other situations where the morphisms between types are somehow internal to the \emph{lifting} package). Hence it is good style to provide explicit rewrite rules how @{const lookup} acts on operations immediately. \ instantiation poly_mapping :: (type, zero) zero begin lift_definition zero_poly_mapping :: "'a \\<^sub>0 'b" is "\k. 0" by simp instance .. end lemma Abs_poly_mapping [simp]: "Abs_poly_mapping (\k. 0) = 0" by (simp add: zero_poly_mapping.abs_eq) lemma lookup_zero [simp]: "lookup 0 k = 0" by transfer rule instantiation poly_mapping :: (type, monoid_add) monoid_add begin lift_definition plus_poly_mapping :: "('a \\<^sub>0 'b) \ ('a \\<^sub>0 'b) \ 'a \\<^sub>0 'b" is "\f1 f2 k. f1 k + f2 k" proof - fix f1 f2 :: "'a \ 'b" assume "finite {k. f1 k \ 0}" and "finite {k. f2 k \ 0}" then have "finite ({k. f1 k \ 0} \ {k. f2 k \ 0})" by auto moreover have "{x. f1 x + f2 x \ 0} \ {k. f1 k \ 0} \ {k. f2 k \ 0}" by auto ultimately show "finite {x. f1 x + f2 x \ 0}" by (blast intro: finite_subset) qed instance by intro_classes (transfer, simp add: fun_eq_iff ac_simps)+ end lemma lookup_add: "lookup (f + g) k = lookup f k + lookup g k" by transfer rule instance poly_mapping :: (type, comm_monoid_add) comm_monoid_add by intro_classes (transfer, simp add: fun_eq_iff ac_simps)+ lemma lookup_sum: "lookup (sum pp X) i = sum (\x. lookup (pp x) i) X" by (induction rule: infinite_finite_induct) (auto simp: lookup_add) (*instance poly_mapping :: (type, "{monoid_add, cancel_semigroup_add}") cancel_semigroup_add by intro_classes (transfer, simp add: fun_eq_iff)+*) instantiation poly_mapping :: (type, cancel_comm_monoid_add) cancel_comm_monoid_add begin lift_definition minus_poly_mapping :: "('a \\<^sub>0 'b) \ ('a \\<^sub>0 'b) \ 'a \\<^sub>0 'b" is "\f1 f2 k. f1 k - f2 k" proof - fix f1 f2 :: "'a \ 'b" assume "finite {k. f1 k \ 0}" and "finite {k. f2 k \ 0}" then have "finite ({k. f1 k \ 0} \ {k. f2 k \ 0})" by auto moreover have "{x. f1 x - f2 x \ 0} \ {k. f1 k \ 0} \ {k. f2 k \ 0}" by auto ultimately show "finite {x. f1 x - f2 x \ 0}" by (blast intro: finite_subset) qed instance by intro_classes (transfer, simp add: fun_eq_iff diff_diff_add)+ end instantiation poly_mapping :: (type, ab_group_add) ab_group_add begin lift_definition uminus_poly_mapping :: "('a \\<^sub>0 'b) \ 'a \\<^sub>0 'b" is uminus by simp instance by intro_classes (transfer, simp add: fun_eq_iff ac_simps)+ end lemma lookup_uminus [simp]: "lookup (- f) k = - lookup f k" by transfer simp lemma lookup_minus: "lookup (f - g) k = lookup f k - lookup g k" by transfer rule subsection \Multiplicative structure\ instantiation poly_mapping :: (zero, zero_neq_one) zero_neq_one begin lift_definition one_poly_mapping :: "'a \\<^sub>0 'b" is "\k. 1 when k = 0" by simp instance by intro_classes (transfer, simp add: fun_eq_iff) end lemma lookup_one: "lookup 1 k = (1 when k = 0)" by transfer rule lemma lookup_one_zero [simp]: "lookup 1 0 = 1" by transfer simp definition prod_fun :: "('a \ 'b) \ ('a \ 'b) \ 'a::monoid_add \ 'b::semiring_0" where "prod_fun f1 f2 k = (\l. f1 l * (\q. (f2 q when k = l + q)))" lemma prod_fun_unfold_prod: fixes f g :: "'a :: monoid_add \ 'b::semiring_0" assumes fin_f: "finite {a. f a \ 0}" assumes fin_g: "finite {b. g b \ 0}" shows "prod_fun f g k = (\(a, b). f a * g b when k = a + b)" proof - let ?C = "{a. f a \ 0} \ {b. g b \ 0}" from fin_f fin_g have "finite ?C" by blast moreover have "{a. \b. (f a * g b when k = a + b) \ 0} \ {b. \a. (f a * g b when k = a + b) \ 0} \ {a. f a \ 0} \ {b. g b \ 0}" by auto ultimately show ?thesis using fin_g by (auto simp add: prod_fun_def Sum_any.cartesian_product [of "{a. f a \ 0} \ {b. g b \ 0}"] Sum_any_right_distrib mult_when) qed lemma finite_prod_fun: fixes f1 f2 :: "'a :: monoid_add \ 'b :: semiring_0" assumes fin1: "finite {l. f1 l \ 0}" and fin2: "finite {q. f2 q \ 0}" shows "finite {k. prod_fun f1 f2 k \ 0}" proof - have *: "finite {k. (\l. f1 l \ 0 \ (\q. f2 q \ 0 \ k = l + q))}" using assms by simp { fix k l have "{q. (f2 q when k = l + q) \ 0} \ {q. f2 q \ 0 \ k = l + q}" by auto with fin2 have "sum f2 {q. f2 q \ 0 \ k = l + q} = (\q. (f2 q when k = l + q))" by (simp add: Sum_any.expand_superset [of "{q. f2 q \ 0 \ k = l + q}"]) } note aux = this have "{k. (\l. f1 l * sum f2 {q. f2 q \ 0 \ k = l + q}) \ 0} \ {k. (\l. f1 l * sum f2 {q. f2 q \ 0 \ k = l + q} \ 0)}" by (auto elim!: Sum_any.not_neutral_obtains_not_neutral) also have "\ \ {k. (\l. f1 l \ 0 \ sum f2 {q. f2 q \ 0 \ k = l + q} \ 0)}" by (auto dest: mult_not_zero) also have "\ \ {k. (\l. f1 l \ 0 \ (\q. f2 q \ 0 \ k = l + q))}" by (auto elim!: sum.not_neutral_contains_not_neutral) finally have "finite {k. (\l. f1 l * sum f2 {q. f2 q \ 0 \ k = l + q}) \ 0}" using * by (rule finite_subset) with aux have "finite {k. (\l. f1 l * (\q. (f2 q when k = l + q))) \ 0}" by simp with fin2 show ?thesis by (simp add: prod_fun_def) qed instantiation poly_mapping :: (monoid_add, semiring_0) semiring_0 begin lift_definition times_poly_mapping :: "('a \\<^sub>0 'b) \ ('a \\<^sub>0 'b) \ 'a \\<^sub>0 'b" is prod_fun by(rule finite_prod_fun) instance proof fix a b c :: "'a \\<^sub>0 'b" show "a * b * c = a * (b * c)" proof transfer fix f g h :: "'a \ 'b" assume fin_f: "finite {a. f a \ 0}" (is "finite ?F") assume fin_g: "finite {b. g b \ 0}" (is "finite ?G") assume fin_h: "finite {c. h c \ 0}" (is "finite ?H") from fin_f fin_g have fin_fg: "finite {(a, b). f a * g b \ 0}" (is "finite ?FG") by (rule finite_mult_not_eq_zero_prodI) from fin_g fin_h have fin_gh: "finite {(b, c). g b * h c \ 0}" (is "finite ?GH") by (rule finite_mult_not_eq_zero_prodI) from fin_f fin_g have fin_fg': "finite {a + b | a b. f a * g b \ 0}" (is "finite ?FG'") by (rule finite_mult_not_eq_zero_sumI) then have fin_fg'': "finite {d. (\(a, b). f a * g b when d = a + b) \ 0}" by (auto intro: finite_Sum_any_not_eq_zero_weakenI) from fin_g fin_h have fin_gh': "finite {b + c | b c. g b * h c \ 0}" (is "finite ?GH'") by (rule finite_mult_not_eq_zero_sumI) then have fin_gh'': "finite {d. (\(b, c). g b * h c when d = b + c) \ 0}" by (auto intro: finite_Sum_any_not_eq_zero_weakenI) show "prod_fun (prod_fun f g) h = prod_fun f (prod_fun g h)" (is "?lhs = ?rhs") proof fix k from fin_f fin_g fin_h fin_fg'' have "?lhs k = (\(ab, c). (\(a, b). f a * g b when ab = a + b) * h c when k = ab + c)" by (simp add: prod_fun_unfold_prod) also have "\ = (\(ab, c). (\(a, b). f a * g b * h c when k = ab + c when ab = a + b))" apply (subst Sum_any_left_distrib) using fin_fg apply (simp add: split_def) apply (subst Sum_any_when_independent [symmetric]) apply (simp add: when_when when_mult mult_when split_def conj_commute) done also have "\ = (\(ab, c, a, b). f a * g b * h c when k = ab + c when ab = a + b)" apply (subst Sum_any.cartesian_product2 [of "(?FG' \ ?H) \ ?FG"]) apply (auto simp add: finite_cartesian_product_iff fin_fg fin_fg' fin_h dest: mult_not_zero) done also have "\ = (\(ab, c, a, b). f a * g b * h c when k = a + b + c when ab = a + b)" by (rule Sum_any.cong) (simp add: split_def when_def) also have "\ = (\(ab, cab). (case cab of (c, a, b) \ f a * g b * h c when k = a + b + c) when ab = (case cab of (c, a, b) \ a + b))" by (simp add: split_def) also have "\ = (\(c, a, b). f a * g b * h c when k = a + b + c)" by (simp add: Sum_any_when_dependent_prod_left) also have "\ = (\(bc, cab). (case cab of (c, a, b) \ f a * g b * h c when k = a + b + c) when bc = (case cab of (c, a, b) \ b + c))" by (simp add: Sum_any_when_dependent_prod_left) also have "\ = (\(bc, c, a, b). f a * g b * h c when k = a + b + c when bc = b + c)" by (simp add: split_def) also have "\ = (\(bc, c, a, b). f a * g b * h c when bc = b + c when k = a + bc)" by (rule Sum_any.cong) (simp add: split_def when_def ac_simps) also have "\ = (\(a, bc, b, c). f a * g b * h c when bc = b + c when k = a + bc)" proof - have "bij (\(a, d, b, c). (d, c, a, b))" by (auto intro!: bijI injI surjI [of _ "\(d, c, a, b). (a, d, b, c)"] simp add: split_def) then show ?thesis by (rule Sum_any.reindex_cong) auto qed also have "\ = (\(a, bc). (\(b, c). f a * g b * h c when bc = b + c when k = a + bc))" apply (subst Sum_any.cartesian_product2 [of "(?F \ ?GH') \ ?GH"]) apply (auto simp add: finite_cartesian_product_iff fin_f fin_gh fin_gh' ac_simps dest: mult_not_zero) done also have "\ = (\(a, bc). f a * (\(b, c). g b * h c when bc = b + c) when k = a + bc)" apply (subst Sum_any_right_distrib) using fin_gh apply (simp add: split_def) apply (subst Sum_any_when_independent [symmetric]) apply (simp add: when_when when_mult mult_when split_def ac_simps) done also from fin_f fin_g fin_h fin_gh'' have "\ = ?rhs k" by (simp add: prod_fun_unfold_prod) finally show "?lhs k = ?rhs k" . qed qed show "(a + b) * c = a * c + b * c" proof transfer fix f g h :: "'a \ 'b" assume fin_f: "finite {k. f k \ 0}" assume fin_g: "finite {k. g k \ 0}" assume fin_h: "finite {k. h k \ 0}" show "prod_fun (\k. f k + g k) h = (\k. prod_fun f h k + prod_fun g h k)" apply (rule ext) apply (auto simp add: prod_fun_def algebra_simps) apply (subst Sum_any.distrib) using fin_f fin_g apply (auto intro: finite_mult_not_eq_zero_rightI) done qed show "a * (b + c) = a * b + a * c" proof transfer fix f g h :: "'a \ 'b" assume fin_f: "finite {k. f k \ 0}" assume fin_g: "finite {k. g k \ 0}" assume fin_h: "finite {k. h k \ 0}" show "prod_fun f (\k. g k + h k) = (\k. prod_fun f g k + prod_fun f h k)" apply (rule ext) apply (auto simp add: prod_fun_def Sum_any.distrib algebra_simps when_add_distrib) apply (subst Sum_any.distrib) apply (simp_all add: algebra_simps) apply (auto intro: fin_g fin_h) apply (subst Sum_any.distrib) apply (simp_all add: algebra_simps) using fin_f apply (rule finite_mult_not_eq_zero_rightI) using fin_f apply (rule finite_mult_not_eq_zero_rightI) done qed show "0 * a = 0" by transfer (simp add: prod_fun_def [abs_def]) show "a * 0 = 0" by transfer (simp add: prod_fun_def [abs_def]) qed end lemma lookup_mult: "lookup (f * g) k = (\l. lookup f l * (\q. lookup g q when k = l + q))" by transfer (simp add: prod_fun_def) instance poly_mapping :: (comm_monoid_add, comm_semiring_0) comm_semiring_0 proof fix a b c :: "'a \\<^sub>0 'b" show "a * b = b * a" proof transfer fix f g :: "'a \ 'b" assume fin_f: "finite {a. f a \ 0}" assume fin_g: "finite {b. g b \ 0}" show "prod_fun f g = prod_fun g f" proof fix k have fin1: "\l. finite {a. (f a when k = l + a) \ 0}" using fin_f by auto have fin2: "\l. finite {b. (g b when k = l + b) \ 0}" using fin_g by auto from fin_f fin_g have "finite {(a, b). f a \ 0 \ g b \ 0}" (is "finite ?AB") by simp show "prod_fun f g k = prod_fun g f k" apply (simp add: prod_fun_def) apply (subst Sum_any_right_distrib) apply (rule fin2) apply (subst Sum_any_right_distrib) apply (rule fin1) apply (subst Sum_any.swap [of ?AB]) apply (fact \finite ?AB\) apply (auto simp add: mult_when ac_simps) done qed qed show "(a + b) * c = a * c + b * c" proof transfer fix f g h :: "'a \ 'b" assume fin_f: "finite {k. f k \ 0}" assume fin_g: "finite {k. g k \ 0}" assume fin_h: "finite {k. h k \ 0}" show "prod_fun (\k. f k + g k) h = (\k. prod_fun f h k + prod_fun g h k)" apply (auto simp add: prod_fun_def fun_eq_iff algebra_simps) apply (subst Sum_any.distrib) using fin_f apply (rule finite_mult_not_eq_zero_rightI) using fin_g apply (rule finite_mult_not_eq_zero_rightI) apply simp_all done qed qed instance poly_mapping :: (monoid_add, semiring_0_cancel) semiring_0_cancel .. instance poly_mapping :: (comm_monoid_add, comm_semiring_0_cancel) comm_semiring_0_cancel .. instance poly_mapping :: (monoid_add, semiring_1) semiring_1 proof fix a :: "'a \\<^sub>0 'b" show "1 * a = a" by transfer (simp add: prod_fun_def [abs_def] when_mult) show "a * 1 = a" apply transfer apply (simp add: prod_fun_def [abs_def] Sum_any_right_distrib Sum_any_left_distrib mult_when) apply (subst when_commute) apply simp done qed instance poly_mapping :: (comm_monoid_add, comm_semiring_1) comm_semiring_1 proof fix a :: "'a \\<^sub>0 'b" show "1 * a = a" by transfer (simp add: prod_fun_def [abs_def]) qed instance poly_mapping :: (monoid_add, semiring_1_cancel) semiring_1_cancel .. instance poly_mapping :: (monoid_add, ring) ring .. instance poly_mapping :: (comm_monoid_add, comm_ring) comm_ring .. instance poly_mapping :: (monoid_add, ring_1) ring_1 .. instance poly_mapping :: (comm_monoid_add, comm_ring_1) comm_ring_1 .. subsection \Single-point mappings\ lift_definition single :: "'a \ 'b \ 'a \\<^sub>0 'b::zero" is "\k v k'. (v when k = k')" by simp lemma inj_single [iff]: "inj (single k)" proof (rule injI, transfer) fix k :: 'b and a b :: "'a::zero" assume "(\k'. a when k = k') = (\k'. b when k = k')" then have "(\k'. a when k = k') k = (\k'. b when k = k') k" by (rule arg_cong) then show "a = b" by simp qed lemma lookup_single: "lookup (single k v) k' = (v when k = k')" by transfer rule lemma lookup_single_eq [simp]: "lookup (single k v) k = v" by transfer simp lemma lookup_single_not_eq: "k \ k' \ lookup (single k v) k' = 0" by transfer simp lemma single_zero [simp]: "single k 0 = 0" by transfer simp lemma single_one [simp]: "single 0 1 = 1" by transfer simp lemma single_add: "single k (a + b) = single k a + single k b" by transfer (simp add: fun_eq_iff when_add_distrib) lemma single_uminus: "single k (- a) = - single k a" by transfer (simp add: fun_eq_iff when_uminus_distrib) lemma single_diff: "single k (a - b) = single k a - single k b" by transfer (simp add: fun_eq_iff when_diff_distrib) lemma single_numeral [simp]: "single 0 (numeral n) = numeral n" by (induct n) (simp_all only: numeral.simps numeral_add single_zero single_one single_add) lemma lookup_numeral: "lookup (numeral n) k = (numeral n when k = 0)" proof - have "lookup (numeral n) k = lookup (single 0 (numeral n)) k" by simp then show ?thesis unfolding lookup_single by simp qed lemma single_of_nat [simp]: "single 0 (of_nat n) = of_nat n" by (induct n) (simp_all add: single_add) lemma lookup_of_nat: "lookup (of_nat n) k = (of_nat n when k = 0)" proof - have "lookup (of_nat n) k = lookup (single 0 (of_nat n)) k" by simp then show ?thesis unfolding lookup_single by simp qed lemma of_nat_single: "of_nat = single 0 \ of_nat" by (simp add: fun_eq_iff) lemma mult_single: "single k a * single l b = single (k + l) (a * b)" proof transfer fix k l :: 'a and a b :: 'b show "prod_fun (\k'. a when k = k') (\k'. b when l = k') = (\k'. a * b when k + l = k')" proof fix k' have "prod_fun (\k'. a when k = k') (\k'. b when l = k') k' = (\n. a * b when l = n when k' = k + n)" by (simp add: prod_fun_def Sum_any_right_distrib mult_when when_mult) also have "\ = (\n. a * b when k' = k + n when l = n)" by (simp add: when_when conj_commute) also have "\ = (a * b when k' = k + l)" by simp also have "\ = (a * b when k + l = k')" by (simp add: when_def) finally show "prod_fun (\k'. a when k = k') (\k'. b when l = k') k' = (\k'. a * b when k + l = k') k'" . qed qed instance poly_mapping :: (monoid_add, semiring_char_0) semiring_char_0 by intro_classes (auto intro: inj_compose inj_of_nat simp add: of_nat_single) instance poly_mapping :: (monoid_add, ring_char_0) ring_char_0 .. lemma single_of_int [simp]: "single 0 (of_int k) = of_int k" by (cases k) (simp_all add: single_diff single_uminus) lemma lookup_of_int: "lookup (of_int l) k = (of_int l when k = 0)" proof - have "lookup (of_int l) k = lookup (single 0 (of_int l)) k" by simp then show ?thesis unfolding lookup_single by simp qed subsection \Integral domains\ instance poly_mapping :: ("{ordered_cancel_comm_monoid_add, linorder}", semiring_no_zero_divisors) semiring_no_zero_divisors text \The @{class "linorder"} constraint is a pragmatic device for the proof --- maybe it can be dropped\ proof fix f g :: "'a \\<^sub>0 'b" assume "f \ 0" and "g \ 0" then show "f * g \ 0" proof transfer fix f g :: "'a \ 'b" define F where "F = {a. f a \ 0}" moreover define G where "G = {a. g a \ 0}" ultimately have [simp]: "\a. f a \ 0 \ a \ F" "\b. g b \ 0 \ b \ G" by simp_all assume "finite {a. f a \ 0}" then have [simp]: "finite F" by simp assume "finite {a. g a \ 0}" then have [simp]: "finite G" by simp assume "f \ (\a. 0)" then obtain a where "f a \ 0" by (auto simp add: fun_eq_iff) assume "g \ (\b. 0)" then obtain b where "g b \ 0" by (auto simp add: fun_eq_iff) from \f a \ 0\ and \g b \ 0\ have "F \ {}" and "G \ {}" by auto note Max_F = \finite F\ \F \ {}\ note Max_G = \finite G\ \G \ {}\ from Max_F and Max_G have [simp]: "Max F \ F" "Max G \ G" by auto from Max_F Max_G have [dest!]: "\a. a \ F \ a \ Max F" "\b. b \ G \ b \ Max G" by auto define q where "q = Max F + Max G" have "(\(a, b). f a * g b when q = a + b) = (\(a, b). f a * g b when q = a + b when a \ F \ b \ G)" by (rule Sum_any.cong) (auto simp add: split_def when_def q_def intro: ccontr) also have "\ = (\(a, b). f a * g b when (Max F, Max G) = (a, b))" proof (rule Sum_any.cong) fix ab :: "'a \ 'a" obtain a b where [simp]: "ab = (a, b)" by (cases ab) simp_all have [dest!]: "a \ Max F \ Max F \ a \ a < Max F" "b \ Max G \ Max G \ b \ b < Max G" by auto show "(case ab of (a, b) \ f a * g b when q = a + b when a \ F \ b \ G) = (case ab of (a, b) \ f a * g b when (Max F, Max G) = (a, b))" by (auto simp add: split_def when_def q_def dest: add_strict_mono [of a "Max F" b "Max G"]) qed also have "\ = (\ab. (case ab of (a, b) \ f a * g b) when (Max F, Max G) = ab)" unfolding split_def when_def by auto also have "\ \ 0" by simp finally have "prod_fun f g q \ 0" by (simp add: prod_fun_unfold_prod) then show "prod_fun f g \ (\k. 0)" by (auto simp add: fun_eq_iff) qed qed instance poly_mapping :: ("{ordered_cancel_comm_monoid_add, linorder}", ring_no_zero_divisors) ring_no_zero_divisors .. instance poly_mapping :: ("{ordered_cancel_comm_monoid_add, linorder}", ring_1_no_zero_divisors) ring_1_no_zero_divisors .. instance poly_mapping :: ("{ordered_cancel_comm_monoid_add, linorder}", idom) idom .. subsection \Mapping order\ instantiation poly_mapping :: (linorder, "{zero, linorder}") linorder begin lift_definition less_poly_mapping :: "('a \\<^sub>0 'b) \ ('a \\<^sub>0 'b) \ bool" is less_fun . lift_definition less_eq_poly_mapping :: "('a \\<^sub>0 'b) \ ('a \\<^sub>0 'b) \ bool" is "\f g. less_fun f g \ f = g" . instance proof (rule class.Orderings.linorder.of_class.intro) show "class.linorder (less_eq :: (_ \\<^sub>0 _) \ _) less" proof (rule linorder_strictI, rule order_strictI) fix f g h :: "'a \\<^sub>0 'b" show "f \ g \ f < g \ f = g" by transfer (rule refl) show "\ f < f" by transfer (rule less_fun_irrefl) show "f < g \ f = g \ g < f" proof transfer fix f g :: "'a \ 'b" assume "finite {k. f k \ 0}" and "finite {k. g k \ 0}" then have "finite ({k. f k \ 0} \ {k. g k \ 0})" by simp moreover have "{k. f k \ g k} \ {k. f k \ 0} \ {k. g k \ 0}" by auto ultimately have "finite {k. f k \ g k}" by (rule rev_finite_subset) then show "less_fun f g \ f = g \ less_fun g f" by (rule less_fun_trichotomy) qed assume "f < g" then show "\ g < f" by transfer (rule less_fun_asym) note \f < g\ moreover assume "g < h" ultimately show "f < h" by transfer (rule less_fun_trans) qed qed end instance poly_mapping :: (linorder, "{ordered_comm_monoid_add, ordered_ab_semigroup_add_imp_le, linorder}") ordered_ab_semigroup_add proof (intro_classes, transfer) fix f g h :: "'a \ 'b" assume *: "less_fun f g \ f = g" { assume "less_fun f g" then obtain k where "f k < g k" "(\k'. k' < k \ f k' = g k')" by (blast elim!: less_funE) then have "h k + f k < h k + g k" "(\k'. k' < k \ h k' + f k' = h k' + g k')" by simp_all then have "less_fun (\k. h k + f k) (\k. h k + g k)" by (blast intro: less_funI) } with * show "less_fun (\k. h k + f k) (\k. h k + g k) \ (\k. h k + f k) = (\k. h k + g k)" by (auto simp add: fun_eq_iff) qed instance poly_mapping :: (linorder, "{ordered_comm_monoid_add, ordered_ab_semigroup_add_imp_le, cancel_comm_monoid_add, linorder}") linordered_cancel_ab_semigroup_add .. instance poly_mapping :: (linorder, "{ordered_comm_monoid_add, ordered_ab_semigroup_add_imp_le, cancel_comm_monoid_add, linorder}") ordered_comm_monoid_add .. instance poly_mapping :: (linorder, "{ordered_comm_monoid_add, ordered_ab_semigroup_add_imp_le, cancel_comm_monoid_add, linorder}") ordered_cancel_comm_monoid_add .. instance poly_mapping :: (linorder, linordered_ab_group_add) linordered_ab_group_add .. text \ For pragmatism we leave out the final elements in the hierarchy: @{class linordered_ring}, @{class linordered_ring_strict}, @{class linordered_idom}; remember that the order instance is a mere technical device, not a deeper algebraic property. \ subsection \Fundamental mapping notions\ lift_definition keys :: "('a \\<^sub>0 'b::zero) \ 'a set" is "\f. {k. f k \ 0}" . lift_definition range :: "('a \\<^sub>0 'b::zero) \ 'b set" is "\f :: 'a \ 'b. Set.range f - {0}" . lemma finite_keys [simp]: "finite (keys f)" by transfer lemma not_in_keys_iff_lookup_eq_zero: "k \ keys f \ lookup f k = 0" by transfer simp lemma lookup_not_eq_zero_eq_in_keys: "lookup f k \ 0 \ k \ keys f" by transfer simp lemma lookup_eq_zero_in_keys_contradict [dest]: "lookup f k = 0 \ \ k \ keys f" by (simp add: not_in_keys_iff_lookup_eq_zero) lemma finite_range [simp]: "finite (Poly_Mapping.range p)" proof transfer fix f :: "'b \ 'a" assume *: "finite {x. f x \ 0}" have "Set.range f - {0} \ f ` {x. f x \ 0}" by auto thus "finite (Set.range f - {0})" by(rule finite_subset)(rule finite_imageI[OF *]) qed lemma in_keys_lookup_in_range [simp]: "k \ keys f \ lookup f k \ range f" by transfer simp lemma in_keys_iff: "x \ (keys s) = (lookup s x \ 0)" by (transfer, simp) lemma keys_zero [simp]: "keys 0 = {}" by transfer simp lemma range_zero [simp]: "range 0 = {}" by transfer auto lemma keys_add: "keys (f + g) \ keys f \ keys g" by transfer auto lemma keys_one [simp]: "keys 1 = {0}" by transfer simp lemma range_one [simp]: "range 1 = {1}" by transfer (auto simp add: when_def) lemma keys_single [simp]: "keys (single k v) = (if v = 0 then {} else {k})" by transfer simp lemma range_single [simp]: "range (single k v) = (if v = 0 then {} else {v})" by transfer (auto simp add: when_def) lemma keys_mult: "keys (f * g) \ {a + b | a b. a \ keys f \ b \ keys g}" apply transfer apply (auto simp add: prod_fun_def dest!: mult_not_zero elim!: Sum_any.not_neutral_obtains_not_neutral) apply blast done lemma setsum_keys_plus_distrib: assumes hom_0: "\k. f k 0 = 0" and hom_plus: "\k. k \ Poly_Mapping.keys p \ Poly_Mapping.keys q \ f k (Poly_Mapping.lookup p k + Poly_Mapping.lookup q k) = f k (Poly_Mapping.lookup p k) + f k (Poly_Mapping.lookup q k)" shows "(\k\Poly_Mapping.keys (p + q). f k (Poly_Mapping.lookup (p + q) k)) = (\k\Poly_Mapping.keys p. f k (Poly_Mapping.lookup p k)) + (\k\Poly_Mapping.keys q. f k (Poly_Mapping.lookup q k))" (is "?lhs = ?p + ?q") proof - let ?A = "Poly_Mapping.keys p \ Poly_Mapping.keys q" have "?lhs = (\k\?A. f k (Poly_Mapping.lookup p k + Poly_Mapping.lookup q k))" apply(rule sum.mono_neutral_cong_left) apply(simp_all add: Poly_Mapping.keys_add) apply(transfer fixing: f) apply(auto simp add: hom_0)[1] apply(transfer fixing: f) apply(auto simp add: hom_0)[1] done also have "\ = (\k\?A. f k (Poly_Mapping.lookup p k) + f k (Poly_Mapping.lookup q k))" by(rule sum.cong)(simp_all add: hom_plus) also have "\ = (\k\?A. f k (Poly_Mapping.lookup p k)) + (\k\?A. f k (Poly_Mapping.lookup q k))" (is "_ = ?p' + ?q'") by(simp add: sum.distrib) also have "?p' = ?p" by (simp add: hom_0 in_keys_iff sum.mono_neutral_cong_right) also have "?q' = ?q" by (simp add: hom_0 in_keys_iff sum.mono_neutral_cong_right) finally show ?thesis . qed subsection \Degree\ definition degree :: "(nat \\<^sub>0 'a::zero) \ nat" where "degree f = Max (insert 0 (Suc ` keys f))" lemma degree_zero [simp]: "degree 0 = 0" unfolding degree_def by transfer simp lemma degree_one [simp]: "degree 1 = 1" unfolding degree_def by transfer simp lemma degree_single_zero [simp]: "degree (single k 0) = 0" unfolding degree_def by transfer simp lemma degree_single_not_zero [simp]: "v \ 0 \ degree (single k v) = Suc k" unfolding degree_def by transfer simp lemma degree_zero_iff [simp]: "degree f = 0 \ f = 0" unfolding degree_def proof transfer fix f :: "nat \ 'a" assume "finite {n. f n \ 0}" then have fin: "finite (insert 0 (Suc ` {n. f n \ 0}))" by auto show "Max (insert 0 (Suc ` {n. f n \ 0})) = 0 \ f = (\n. 0)" (is "?P \ ?Q") proof assume ?P have "{n. f n \ 0} = {}" proof (rule ccontr) assume "{n. f n \ 0} \ {}" then obtain n where "n \ {n. f n \ 0}" by blast then have "{n. f n \ 0} = insert n {n. f n \ 0}" by auto then have "Suc ` {n. f n \ 0} = insert (Suc n) (Suc ` {n. f n \ 0})" by auto with \?P\ have "Max (insert 0 (insert (Suc n) (Suc ` {n. f n \ 0}))) = 0" by simp then have "Max (insert (Suc n) (insert 0 (Suc ` {n. f n \ 0}))) = 0" by (simp add: insert_commute) with fin have "max (Suc n) (Max (insert 0 (Suc ` {n. f n \ 0}))) = 0" by simp then show False by simp qed then show ?Q by (simp add: fun_eq_iff) next assume ?Q then show ?P by simp qed qed lemma degree_greater_zero_in_keys: assumes "0 < degree f" shows "degree f - 1 \ keys f" proof - from assms have "keys f \ {}" by (auto simp add: degree_def) then show ?thesis unfolding degree_def by (simp add: mono_Max_commute [symmetric] mono_Suc) qed lemma in_keys_less_degree: "n \ keys f \ n < degree f" unfolding degree_def by transfer (auto simp add: Max_gr_iff) lemma beyond_degree_lookup_zero: "degree f \ n \ lookup f n = 0" unfolding degree_def by transfer auto lemma degree_add: "degree (f + g) \ max (degree f) (Poly_Mapping.degree g)" unfolding degree_def proof transfer fix f g :: "nat \ 'a" assume f: "finite {x. f x \ 0}" assume g: "finite {x. g x \ 0}" let ?f = "Max (insert 0 (Suc ` {k. f k \ 0}))" let ?g = "Max (insert 0 (Suc ` {k. g k \ 0}))" have "Max (insert 0 (Suc ` {k. f k + g k \ 0})) \ Max (insert 0 (Suc ` ({k. f k \ 0} \ {k. g k \ 0})))" by (rule Max.subset_imp) (insert f g, auto) also have "\ = max ?f ?g" using f g by (simp_all add: image_Un Max_Un [symmetric]) finally show "Max (insert 0 (Suc ` {k. f k + g k \ 0})) \ max (Max (insert 0 (Suc ` {k. f k \ 0}))) (Max (insert 0 (Suc ` {k. g k \ 0})))" . qed lemma sorted_list_of_set_keys: "sorted_list_of_set (keys f) = filter (\k. k \ keys f) [0..Inductive structure\ lift_definition update :: "'a \ 'b \ ('a \\<^sub>0 'b::zero) \ 'a \\<^sub>0 'b" is "\k v f. f(k := v)" proof - fix f :: "'a \ 'b" and k' v assume "finite {k. f k \ 0}" then have "finite (insert k' {k. f k \ 0})" by simp then show "finite {k. (f(k' := v)) k \ 0}" by (rule rev_finite_subset) auto qed lemma update_induct [case_names const update]: assumes const': "P 0" assumes update': "\f a b. a \ keys f \ b \ 0 \ P f \ P (update a b f)" shows "P f" proof - obtain g where "f = Abs_poly_mapping g" and "finite {a. g a \ 0}" by (cases f) simp_all define Q where "Q g = P (Abs_poly_mapping g)" for g from \finite {a. g a \ 0}\ have "Q g" proof (induct g rule: finite_update_induct) case const with const' Q_def show ?case by simp next case (update a b g) from \finite {a. g a \ 0}\ \g a = 0\ have "a \ keys (Abs_poly_mapping g)" by (simp add: Abs_poly_mapping_inverse keys.rep_eq) moreover note \b \ 0\ moreover from \Q g\ have "P (Abs_poly_mapping g)" by (simp add: Q_def) ultimately have "P (update a b (Abs_poly_mapping g))" by (rule update') also from \finite {a. g a \ 0}\ have "update a b (Abs_poly_mapping g) = Abs_poly_mapping (g(a := b))" by (simp add: update.abs_eq eq_onp_same_args) finally show ?case by (simp add: Q_def fun_upd_def) qed then show ?thesis by (simp add: Q_def \f = Abs_poly_mapping g\) qed lemma lookup_update: "lookup (update k v f) k' = (if k = k' then v else lookup f k')" by transfer simp lemma keys_update: "keys (update k v f) = (if v = 0 then keys f - {k} else insert k (keys f))" by transfer auto subsection \Quasi-functorial structure\ lift_definition map :: "('b::zero \ 'c::zero) \ ('a \\<^sub>0 'b) \ ('a \\<^sub>0 'c::zero)" is "\g f k. g (f k) when f k \ 0" by simp context fixes f :: "'b \ 'a" assumes inj_f: "inj f" begin lift_definition map_key :: "('a \\<^sub>0 'c::zero) \ 'b \\<^sub>0 'c" is "\p. p \ f" proof - fix g :: "'c \ 'd" and p :: "'a \ 'c" assume "finite {x. p x \ 0}" hence "finite (f ` {y. p (f y) \ 0})" by(rule finite_subset[rotated]) auto thus "finite {x. (p \ f) x \ 0}" unfolding o_def by(rule finite_imageD)(rule subset_inj_on[OF inj_f], simp) qed end lemma map_key_compose: assumes [transfer_rule]: "inj f" "inj g" shows "map_key f (map_key g p) = map_key (g \ f) p" proof - from assms have [transfer_rule]: "inj (g \ f)" by(simp add: inj_compose) show ?thesis by transfer(simp add: o_assoc) qed lemma map_key_id: "map_key (\x. x) p = p" proof - have [transfer_rule]: "inj (\x. x)" by simp show ?thesis by transfer(simp add: o_def) qed context fixes f :: "'a \ 'b" assumes inj_f [transfer_rule]: "inj f" begin lemma map_key_map: "map_key f (map g p) = map g (map_key f p)" by transfer (simp add: fun_eq_iff) lemma map_key_plus: "map_key f (p + q) = map_key f p + map_key f q" by transfer (simp add: fun_eq_iff) lemma keys_map_key: "keys (map_key f p) = f -` keys p" by transfer auto lemma map_key_zero [simp]: "map_key f 0 = 0" by transfer (simp add: fun_eq_iff) lemma map_key_single [simp]: "map_key f (single (f k) v) = single k v" by transfer (simp add: fun_eq_iff inj_onD [OF inj_f] when_def) end lemma mult_map_scale_conv_mult: "map ((*) s) p = single 0 s * p" proof(transfer fixing: s) fix p :: "'a \ 'b" assume *: "finite {x. p x \ 0}" { fix x have "prod_fun (\k'. s when 0 = k') p x = (\l :: 'a. if l = 0 then s * (\q. p q when x = q) else 0)" by(auto simp add: prod_fun_def when_def intro: Sum_any.cong simp del: Sum_any.delta) also have "\ = (\k. s * p k when p k \ 0) x" by(simp add: when_def) also note calculation } then show "(\k. s * p k when p k \ 0) = prod_fun (\k'. s when 0 = k') p" by(simp add: fun_eq_iff) qed lemma map_single [simp]: "(c = 0 \ f 0 = 0) \ map f (single x c) = single x (f c)" by transfer(auto simp add: fun_eq_iff when_def) lemma map_eq_zero_iff: "map f p = 0 \ (\k \ keys p. f (lookup p k) = 0)" by transfer(auto simp add: fun_eq_iff when_def) subsection \Canonical dense representation of @{typ "nat \\<^sub>0 'a"}\ abbreviation no_trailing_zeros :: "'a :: zero list \ bool" where "no_trailing_zeros \ no_trailing ((=) 0)" lift_definition "nth" :: "'a list \ (nat \\<^sub>0 'a::zero)" is "nth_default 0" by (fact finite_nth_default_neq_default) text \ The opposite direction is directly specified on (later) type \nat_mapping\. \ lemma nth_Nil [simp]: "nth [] = 0" by transfer (simp add: fun_eq_iff) lemma nth_singleton [simp]: "nth [v] = single 0 v" proof (transfer, rule ext) fix n :: nat and v :: 'a show "nth_default 0 [v] n = (v when 0 = n)" by (auto simp add: nth_default_def nth_append) qed lemma nth_replicate [simp]: "nth (replicate n 0 @ [v]) = single n v" proof (transfer, rule ext) fix m n :: nat and v :: 'a show "nth_default 0 (replicate n 0 @ [v]) m = (v when n = m)" by (auto simp add: nth_default_def nth_append) qed lemma nth_strip_while [simp]: "nth (strip_while ((=) 0) xs) = nth xs" by transfer (fact nth_default_strip_while_dflt) lemma nth_strip_while' [simp]: "nth (strip_while (\k. k = 0) xs) = nth xs" by (subst eq_commute) (fact nth_strip_while) lemma nth_eq_iff: "nth xs = nth ys \ strip_while (HOL.eq 0) xs = strip_while (HOL.eq 0) ys" by transfer (simp add: nth_default_eq_iff) lemma lookup_nth [simp]: "lookup (nth xs) = nth_default 0 xs" by (fact nth.rep_eq) lemma keys_nth [simp]: "keys (nth xs) = fst ` {(n, v) \ set (enumerate 0 xs). v \ 0}" proof transfer fix xs :: "'a list" { fix n assume "nth_default 0 xs n \ 0" then have "n < length xs" and "xs ! n \ 0" by (auto simp add: nth_default_def split: if_splits) then have "(n, xs ! n) \ {(n, v). (n, v) \ set (enumerate 0 xs) \ v \ 0}" (is "?x \ ?A") by (auto simp add: in_set_conv_nth enumerate_eq_zip) then have "fst ?x \ fst ` ?A" by blast then have "n \ fst ` {(n, v). (n, v) \ set (enumerate 0 xs) \ v \ 0}" by simp } then show "{k. nth_default 0 xs k \ 0} = fst ` {(n, v). (n, v) \ set (enumerate 0 xs) \ v \ 0}" by (auto simp add: in_enumerate_iff_nth_default_eq) qed lemma range_nth [simp]: "range (nth xs) = set xs - {0}" by transfer simp lemma degree_nth: "no_trailing_zeros xs \ degree (nth xs) = length xs" unfolding degree_def proof transfer fix xs :: "'a list" assume *: "no_trailing_zeros xs" let ?A = "{n. nth_default 0 xs n \ 0}" let ?f = "nth_default 0 xs" let ?bound = "Max (insert 0 (Suc ` {n. ?f n \ 0}))" show "?bound = length xs" proof (cases "xs = []") case False with * obtain n where n: "n < length xs" "xs ! n \ 0" by (fastforce simp add: no_trailing_unfold last_conv_nth neq_Nil_conv) then have "?bound = Max (Suc ` {k. (k < length xs \ xs ! k \ (0::'a)) \ k < length xs})" by (subst Max_insert) (auto simp add: nth_default_def) also let ?A = "{k. k < length xs \ xs ! k \ 0}" have "{k. (k < length xs \ xs ! k \ (0::'a)) \ k < length xs} = ?A" by auto also have "Max (Suc ` ?A) = Suc (Max ?A)" using n by (subst mono_Max_commute [where f = Suc, symmetric]) (auto simp add: mono_Suc) also { have "Max ?A \ ?A" using n Max_in [of ?A] by fastforce hence "Suc (Max ?A) \ length xs" by simp moreover from * False have "length xs - 1 \ ?A" by(auto simp add: no_trailing_unfold last_conv_nth) hence "length xs - 1 \ Max ?A" using Max_ge[of ?A "length xs - 1"] by auto hence "length xs \ Suc (Max ?A)" by simp ultimately have "Suc (Max ?A) = length xs" by simp } finally show ?thesis . qed simp qed lemma nth_trailing_zeros [simp]: "nth (xs @ replicate n 0) = nth xs" by transfer simp lemma nth_idem: "nth (List.map (lookup f) [0.. n" shows "nth (List.map (lookup f) [0..Canonical sparse representation of @{typ "'a \\<^sub>0 'b"}\ lift_definition the_value :: "('a \ 'b) list \ 'a \\<^sub>0 'b::zero" is "\xs k. case map_of xs k of None \ 0 | Some v \ v" proof - fix xs :: "('a \ 'b) list" have fin: "finite {k. \v. map_of xs k = Some v}" using finite_dom_map_of [of xs] unfolding dom_def by auto then show "finite {k. (case map_of xs k of None \ 0 | Some v \ v) \ 0}" using fin by (simp split: option.split) qed definition items :: "('a::linorder \\<^sub>0 'b::zero) \ ('a \ 'b) list" where "items f = List.map (\k. (k, lookup f k)) (sorted_list_of_set (keys f))" text \ For the canonical sparse representation we provide both directions of morphisms since the specification of ordered association lists in theory \OAList\ will support arbitrary linear orders @{class linorder} as keys, not just natural numbers @{typ nat}. \ lemma the_value_items [simp]: "the_value (items f) = f" unfolding items_def by transfer (simp add: fun_eq_iff map_of_map_restrict restrict_map_def) lemma lookup_the_value: "lookup (the_value xs) k = (case map_of xs k of None \ 0 | Some v \ v)" by transfer rule lemma items_the_value: assumes "sorted (List.map fst xs)" and "distinct (List.map fst xs)" and "0 \ snd ` set xs" shows "items (the_value xs) = xs" proof - from assms have "sorted_list_of_set (set (List.map fst xs)) = List.map fst xs" unfolding sorted_list_of_set_sort_remdups by (simp add: distinct_remdups_id sorted_sort_id) moreover from assms have "keys (the_value xs) = fst ` set xs" by transfer (auto simp add: image_def split: option.split dest: set_map_of_compr) ultimately show ?thesis unfolding items_def using assms by (auto simp add: lookup_the_value intro: map_idI) qed lemma the_value_Nil [simp]: "the_value [] = 0" by transfer (simp add: fun_eq_iff) lemma the_value_Cons [simp]: "the_value (x # xs) = update (fst x) (snd x) (the_value xs)" by transfer (simp add: fun_eq_iff) lemma items_zero [simp]: "items 0 = []" unfolding items_def by simp lemma items_one [simp]: "items 1 = [(0, 1)]" unfolding items_def by transfer simp lemma items_single [simp]: "items (single k v) = (if v = 0 then [] else [(k, v)])" unfolding items_def by simp lemma in_set_items_iff [simp]: "(k, v) \ set (items f) \ k \ keys f \ lookup f k = v" unfolding items_def by transfer auto subsection \Size estimation\ context fixes f :: "'a \ nat" and g :: "'b :: zero \ nat" begin definition poly_mapping_size :: "('a \\<^sub>0 'b) \ nat" where "poly_mapping_size m = g 0 + (\k \ keys m. Suc (f k + g (lookup m k)))" lemma poly_mapping_size_0 [simp]: "poly_mapping_size 0 = g 0" by (simp add: poly_mapping_size_def) lemma poly_mapping_size_single [simp]: "poly_mapping_size (single k v) = (if v = 0 then g 0 else g 0 + f k + g v + 1)" unfolding poly_mapping_size_def by transfer simp lemma keys_less_poly_mapping_size: "k \ keys m \ f k + g (lookup m k) < poly_mapping_size m" unfolding poly_mapping_size_def proof transfer fix k :: 'a and m :: "'a \ 'b" and f :: "'a \ nat" and g let ?keys = "{k. m k \ 0}" assume *: "finite ?keys" "k \ ?keys" then have "f k + g (m k) = (\k' \ ?keys. f k' + g (m k') when k' = k)" by (simp add: sum.delta when_def) also have "\ < (\k' \ ?keys. Suc (f k' + g (m k')))" using * by (intro sum_strict_mono) (auto simp add: when_def) also have "\ \ g 0 + \" by simp finally have "f k + g (m k) < \" . then show "f k + g (m k) < g 0 + (\k | m k \ 0. Suc (f k + g (m k)))" by simp qed lemma lookup_le_poly_mapping_size: "g (lookup m k) \ poly_mapping_size m" proof (cases "k \ keys m") case True with keys_less_poly_mapping_size [of k m] show ?thesis by simp next case False then show ?thesis by (simp add: Poly_Mapping.poly_mapping_size_def in_keys_iff) qed lemma poly_mapping_size_estimation: "k \ keys m \ y \ f k + g (lookup m k) \ y < poly_mapping_size m" using keys_less_poly_mapping_size by (auto intro: le_less_trans) lemma poly_mapping_size_estimation2: assumes "v \ range m" and "y \ g v" shows "y < poly_mapping_size m" proof - from assms obtain k where *: "lookup m k = v" "v \ 0" by transfer blast from * have "k \ keys m" by (simp add: in_keys_iff) then show ?thesis proof (rule poly_mapping_size_estimation) from assms * have "y \ g (lookup m k)" by simp then show "y \ f k + g (lookup m k)" by simp qed qed end lemma poly_mapping_size_one [simp]: "poly_mapping_size f g 1 = g 0 + f 0 + g 1 + 1" unfolding poly_mapping_size_def by transfer simp lemma poly_mapping_size_cong [fundef_cong]: "m = m' \ g 0 = g' 0 \ (\k. k \ keys m' \ f k = f' k) \ (\v. v \ range m' \ g v = g' v) \ poly_mapping_size f g m = poly_mapping_size f' g' m'" by (auto simp add: poly_mapping_size_def intro!: sum.cong) instantiation poly_mapping :: (type, zero) size begin definition "size = poly_mapping_size (\_. 0) (\_. 0)" instance .. end subsection \Further mapping operations and properties\ text \It is like in algebra: there are many definitions, some are also used\ lift_definition mapp :: "('a \ 'b :: zero \ 'c :: zero) \ ('a \\<^sub>0 'b) \ ('a \\<^sub>0 'c)" is "\f p k. (if k \ keys p then f k (lookup p k) else 0)" by simp lemma mapp_cong [fundef_cong]: "\ m = m'; \k. k \ keys m' \ f k (lookup m' k) = f' k (lookup m' k) \ \ mapp f m = mapp f' m'" by transfer (auto simp add: fun_eq_iff) lemma lookup_mapp: "lookup (mapp f p) k = (f k (lookup p k) when k \ keys p)" by (simp add: mapp.rep_eq) lemma keys_mapp_subset: "keys (mapp f p) \ keys p" by (meson in_keys_iff mapp.rep_eq subsetI) subsection\Free Abelian Groups Over a Type\ abbreviation frag_of :: "'a \ 'a \\<^sub>0 int" where "frag_of c \ Poly_Mapping.single c (1::int)" lemma lookup_frag_of [simp]: "Poly_Mapping.lookup(frag_of c) = (\x. if x = c then 1 else 0)" by (force simp add: lookup_single_not_eq) lemma frag_of_nonzero [simp]: "frag_of a \ 0" proof - let ?f = "\x. if x = a then 1 else (0::int)" have "?f \ (\x. 0::int)" by (auto simp: fun_eq_iff) then have "Poly_Mapping.lookup (Abs_poly_mapping ?f) \ Poly_Mapping.lookup (Abs_poly_mapping (\x. 0))" by fastforce then show ?thesis by (metis lookup_single_eq lookup_zero) qed definition frag_cmul :: "int \ ('a \\<^sub>0 int) \ ('a \\<^sub>0 int)" where "frag_cmul c a = Abs_poly_mapping (\x. c * Poly_Mapping.lookup a x)" lemma frag_cmul_zero [simp]: "frag_cmul 0 x = 0" by (simp add: frag_cmul_def) lemma frag_cmul_zero2 [simp]: "frag_cmul c 0 = 0" by (simp add: frag_cmul_def) lemma frag_cmul_one [simp]: "frag_cmul 1 x = x" by (auto simp: frag_cmul_def Poly_Mapping.poly_mapping.lookup_inverse) lemma frag_cmul_minus_one [simp]: "frag_cmul (-1) x = -x" by (simp add: frag_cmul_def uminus_poly_mapping_def poly_mapping_eqI) lemma frag_cmul_cmul [simp]: "frag_cmul c (frag_cmul d x) = frag_cmul (c*d) x" by (simp add: frag_cmul_def mult_ac) lemma lookup_frag_cmul [simp]: "poly_mapping.lookup (frag_cmul c x) i = c * poly_mapping.lookup x i" by (simp add: frag_cmul_def) lemma minus_frag_cmul [simp]: "- frag_cmul k x = frag_cmul (-k) x" by (simp add: poly_mapping_eqI) lemma keys_frag_of: "Poly_Mapping.keys(frag_of a) = {a}" by simp lemma finite_cmul_nonzero: "finite {x. c * Poly_Mapping.lookup a x \ (0::int)}" by simp lemma keys_cmul: "Poly_Mapping.keys(frag_cmul c a) \ Poly_Mapping.keys a" using finite_cmul_nonzero [of c a] by (metis lookup_frag_cmul mult_zero_right not_in_keys_iff_lookup_eq_zero subsetI) lemma keys_cmul_iff [iff]: "i \ Poly_Mapping.keys (frag_cmul c x) \ i \ Poly_Mapping.keys x \ c \ 0" - apply (auto simp: ) + apply auto apply (meson subsetD keys_cmul) by (metis in_keys_iff lookup_frag_cmul mult_eq_0_iff) lemma keys_minus [simp]: "Poly_Mapping.keys(-a) = Poly_Mapping.keys a" by (metis (no_types, opaque_lifting) in_keys_iff lookup_uminus neg_equal_0_iff_equal subsetI subset_antisym) lemma keys_diff: "Poly_Mapping.keys(a - b) \ Poly_Mapping.keys a \ Poly_Mapping.keys b" by (auto simp add: in_keys_iff lookup_minus) lemma keys_eq_empty [simp]: "Poly_Mapping.keys c = {} \ c = 0" by (metis in_keys_iff keys_zero lookup_zero poly_mapping_eqI) lemma frag_cmul_eq_0_iff [simp]: "frag_cmul k c = 0 \ k=0 \ c=0" by auto (metis subsetI subset_antisym keys_cmul_iff keys_eq_empty) lemma frag_of_eq: "frag_of x = frag_of y \ x = y" by (metis lookup_single_eq lookup_single_not_eq zero_neq_one) lemma frag_cmul_distrib: "frag_cmul (c+d) a = frag_cmul c a + frag_cmul d a" by (simp add: frag_cmul_def plus_poly_mapping_def int_distrib) lemma frag_cmul_distrib2: "frag_cmul c (a+b) = frag_cmul c a + frag_cmul c b" proof - have "finite {x. poly_mapping.lookup a x + poly_mapping.lookup b x \ 0}" using keys_add [of a b] by (metis (no_types, lifting) finite_keys finite_subset keys.rep_eq lookup_add mem_Collect_eq subsetI) then show ?thesis by (simp add: frag_cmul_def plus_poly_mapping_def int_distrib) qed lemma frag_cmul_diff_distrib: "frag_cmul (a - b) c = frag_cmul a c - frag_cmul b c" by (auto simp: left_diff_distrib lookup_minus poly_mapping_eqI) lemma frag_cmul_sum: "frag_cmul a (sum b I) = (\i\I. frag_cmul a (b i))" proof (induction rule: infinite_finite_induct) case (insert i I) then show ?case by (auto simp: algebra_simps frag_cmul_distrib2) qed auto lemma keys_sum: "Poly_Mapping.keys(sum b I) \ (\i \I. Poly_Mapping.keys(b i))" proof (induction I rule: infinite_finite_induct) case (insert i I) then show ?case using keys_add [of "b i" "sum b I"] by auto qed auto definition frag_extend :: "('b \ 'a \\<^sub>0 int) \ ('b \\<^sub>0 int) \ 'a \\<^sub>0 int" where "frag_extend b x \ (\i \ Poly_Mapping.keys x. frag_cmul (Poly_Mapping.lookup x i) (b i))" lemma frag_extend_0 [simp]: "frag_extend b 0 = 0" by (simp add: frag_extend_def) lemma frag_extend_of [simp]: "frag_extend f (frag_of a) = f a" by (simp add: frag_extend_def) lemma frag_extend_cmul: "frag_extend f (frag_cmul c x) = frag_cmul c (frag_extend f x)" by (auto simp: frag_extend_def frag_cmul_sum intro: sum.mono_neutral_cong_left) lemma frag_extend_minus: "frag_extend f (- x) = - (frag_extend f x)" using frag_extend_cmul [of f "-1"] by simp lemma frag_extend_add: "frag_extend f (a+b) = (frag_extend f a) + (frag_extend f b)" proof - have *: "(\i\Poly_Mapping.keys a. frag_cmul (poly_mapping.lookup a i) (f i)) = (\i\Poly_Mapping.keys a \ Poly_Mapping.keys b. frag_cmul (poly_mapping.lookup a i) (f i))" "(\i\Poly_Mapping.keys b. frag_cmul (poly_mapping.lookup b i) (f i)) = (\i\Poly_Mapping.keys a \ Poly_Mapping.keys b. frag_cmul (poly_mapping.lookup b i) (f i))" by (auto simp: in_keys_iff intro: sum.mono_neutral_cong_left) have "frag_extend f (a+b) = (\i\Poly_Mapping.keys (a + b). frag_cmul (poly_mapping.lookup a i) (f i) + frag_cmul (poly_mapping.lookup b i) (f i)) " by (auto simp: frag_extend_def Poly_Mapping.lookup_add frag_cmul_distrib) also have "... = (\i \ Poly_Mapping.keys a \ Poly_Mapping.keys b. frag_cmul (poly_mapping.lookup a i) (f i) + frag_cmul (poly_mapping.lookup b i) (f i))" apply (rule sum.mono_neutral_cong_left) using keys_add [of a b] apply (auto simp add: in_keys_iff plus_poly_mapping.rep_eq frag_cmul_distrib [symmetric]) done also have "... = (frag_extend f a) + (frag_extend f b)" by (auto simp: * sum.distrib frag_extend_def) finally show ?thesis . qed lemma frag_extend_diff: "frag_extend f (a-b) = (frag_extend f a) - (frag_extend f b)" by (metis (no_types, opaque_lifting) add_uminus_conv_diff frag_extend_add frag_extend_minus) lemma frag_extend_sum: "finite I \ frag_extend f (\i\I. g i) = sum (frag_extend f o g) I" by (induction I rule: finite_induct) (simp_all add: frag_extend_add) lemma frag_extend_eq: "(\f. f \ Poly_Mapping.keys c \ g f = h f) \ frag_extend g c = frag_extend h c" by (simp add: frag_extend_def) lemma frag_extend_eq_0: "(\x. x \ Poly_Mapping.keys c \ f x = 0) \ frag_extend f c = 0" by (simp add: frag_extend_def) lemma keys_frag_extend: "Poly_Mapping.keys(frag_extend f c) \ (\x \ Poly_Mapping.keys c. Poly_Mapping.keys(f x))" unfolding frag_extend_def using keys_sum by fastforce lemma frag_expansion: "a = frag_extend frag_of a" proof - have *: "finite I \ Poly_Mapping.lookup (\i\I. frag_cmul (Poly_Mapping.lookup a i) (frag_of i)) j = (if j \ I then Poly_Mapping.lookup a j else 0)" for I j by (induction I rule: finite_induct) (auto simp: lookup_single lookup_add) show ?thesis unfolding frag_extend_def by (rule poly_mapping_eqI) (fastforce simp add: in_keys_iff *) qed lemma frag_closure_minus_cmul: assumes "P 0" and P: "\x y. \P x; P y\ \ P(x - y)" "P c" shows "P(frag_cmul k c)" proof - have "P (frag_cmul (int n) c)" for n apply (induction n) apply (simp_all add: assms frag_cmul_distrib) by (metis add.left_neutral add_diff_cancel_right' add_uminus_conv_diff P) then show ?thesis by (metis (no_types, opaque_lifting) add_diff_eq assms(2) diff_add_cancel frag_cmul_distrib int_diff_cases) qed lemma frag_induction [consumes 1, case_names zero one diff]: assumes supp: "Poly_Mapping.keys c \ S" and 0: "P 0" and sing: "\x. x \ S \ P(frag_of x)" and diff: "\a b. \P a; P b\ \ P(a - b)" shows "P c" proof - have "P (\i\I. frag_cmul (poly_mapping.lookup c i) (frag_of i))" if "I \ Poly_Mapping.keys c" for I using finite_subset [OF that finite_keys [of c]] that supp proof (induction I arbitrary: c rule: finite_induct) case empty then show ?case by (auto simp: 0) next case (insert i I c) have ab: "a+b = a - (0 - b)" for a b :: "'a \\<^sub>0 int" by simp have Pfrag: "P (frag_cmul (poly_mapping.lookup c i) (frag_of i))" by (metis "0" diff frag_closure_minus_cmul insert.prems insert_subset sing subset_iff) show ?case apply (simp add: insert.hyps) apply (subst ab) using insert apply (blast intro: assms Pfrag) done qed then show ?thesis by (subst frag_expansion) (auto simp add: frag_extend_def) qed lemma frag_extend_compose: "frag_extend f (frag_extend (frag_of o g) c) = frag_extend (f o g) c" using subset_UNIV by (induction c rule: frag_induction) (auto simp: frag_extend_diff) lemma frag_split: fixes c :: "'a \\<^sub>0 int" assumes "Poly_Mapping.keys c \ S \ T" obtains d e where "Poly_Mapping.keys d \ S" "Poly_Mapping.keys e \ T" "d + e = c" proof let ?d = "frag_extend (\f. if f \ S then frag_of f else 0) c" let ?e = "frag_extend (\f. if f \ S then 0 else frag_of f) c" show "Poly_Mapping.keys ?d \ S" "Poly_Mapping.keys ?e \ T" using assms by (auto intro!: order_trans [OF keys_frag_extend] split: if_split_asm) show "?d + ?e = c" using assms proof (induction c rule: frag_induction) case (diff a b) then show ?case by (metis (no_types, lifting) frag_extend_diff add_diff_eq diff_add_eq diff_add_eq_diff_diff_swap) qed auto qed hide_const (open) lookup single update keys range map map_key degree nth the_value items foldr mapp end \ No newline at end of file diff --git a/src/HOL/Probability/Distributions.thy b/src/HOL/Probability/Distributions.thy --- a/src/HOL/Probability/Distributions.thy +++ b/src/HOL/Probability/Distributions.thy @@ -1,1378 +1,1378 @@ (* Title: HOL/Probability/Distributions.thy Author: Sudeep Kanav, TU München Author: Johannes Hölzl, TU München Author: Jeremy Avigad, CMU *) section \Properties of Various Distributions\ theory Distributions imports Convolution Information begin lemma (in prob_space) distributed_affine: fixes f :: "real \ ennreal" assumes f: "distributed M lborel X f" assumes c: "c \ 0" shows "distributed M lborel (\x. t + c * X x) (\x. f ((x - t) / c) / \c\)" unfolding distributed_def proof safe have [measurable]: "f \ borel_measurable borel" using f by (simp add: distributed_def) have [measurable]: "X \ borel_measurable M" using f by (simp add: distributed_def) show "(\x. f ((x - t) / c) / \c\) \ borel_measurable lborel" by simp show "random_variable lborel (\x. t + c * X x)" by simp have eq: "ennreal \c\ * (f x / ennreal \c\) = f x" for x using c by (cases "f x") (auto simp: divide_ennreal ennreal_mult[symmetric] ennreal_top_divide ennreal_mult_top) have "density lborel f = distr M lborel X" using f by (simp add: distributed_def) with c show "distr M lborel (\x. t + c * X x) = density lborel (\x. f ((x - t) / c) / ennreal \c\)" by (subst (2) lborel_real_affine[where c="c" and t="t"]) (simp_all add: density_density_eq density_distr distr_distr field_simps eq cong: distr_cong) qed lemma (in prob_space) distributed_affineI: fixes f :: "real \ ennreal" and c :: real assumes f: "distributed M lborel (\x. (X x - t) / c) (\x. \c\ * f (x * c + t))" assumes c: "c \ 0" shows "distributed M lborel X f" proof - have eq: "f x * ennreal \c\ / ennreal \c\ = f x" for x using c by (simp add: ennreal_times_divide[symmetric]) show ?thesis using distributed_affine[OF f c, where t=t] c by (simp add: field_simps eq) qed lemma (in prob_space) distributed_AE2: assumes [measurable]: "distributed M N X f" "Measurable.pred N P" shows "(AE x in M. P (X x)) \ (AE x in N. 0 < f x \ P x)" proof - have "(AE x in M. P (X x)) \ (AE x in distr M N X. P x)" by (simp add: AE_distr_iff) also have "\ \ (AE x in density N f. P x)" unfolding distributed_distr_eq_density[OF assms(1)] .. also have "\ \ (AE x in N. 0 < f x \ P x)" by (rule AE_density) simp finally show ?thesis . qed subsection \Erlang\ lemma nn_intergal_power_times_exp_Icc: assumes [arith]: "0 \ a" shows "(\\<^sup>+x. ennreal (x^k * exp (-x)) * indicator {0 .. a} x \lborel) = (1 - (\n\k. (a^n * exp (-a)) / fact n)) * fact k" (is "?I = _") proof - let ?f = "\k x. x^k * exp (-x) / fact k" let ?F = "\k x. - (\n\k. (x^n * exp (-x)) / fact n)" have "?I * (inverse (real_of_nat (fact k))) = (\\<^sup>+x. ennreal (x^k * exp (-x)) * indicator {0 .. a} x * (inverse (real_of_nat (fact k))) \lborel)" by (intro nn_integral_multc[symmetric]) auto also have "\ = (\\<^sup>+x. ennreal (?f k x) * indicator {0 .. a} x \lborel)" by (intro nn_integral_cong) (simp add: field_simps ennreal_mult'[symmetric] indicator_mult_ennreal) also have "\ = ennreal (?F k a - ?F k 0)" proof (rule nn_integral_FTC_Icc) fix x assume "x \ {0..a}" show "DERIV (?F k) x :> ?f k x" proof(induction k) case 0 show ?case by (auto intro!: derivative_eq_intros) next case (Suc k) have "DERIV (\x. ?F k x - (x^Suc k * exp (-x)) / fact (Suc k)) x :> ?f k x - ((real (Suc k) - x) * x ^ k * exp (- x)) / (fact (Suc k))" by (intro DERIV_diff Suc) (auto intro!: derivative_eq_intros simp del: fact_Suc power_Suc simp add: field_simps power_Suc[symmetric]) also have "(\x. ?F k x - (x^Suc k * exp (-x)) / fact (Suc k)) = ?F (Suc k)" by simp also have "?f k x - ((real (Suc k) - x) * x ^ k * exp (- x)) / (fact (Suc k)) = ?f (Suc k) x" by (auto simp: field_simps simp del: fact_Suc) (simp_all add: of_nat_Suc field_simps) finally show ?case . qed qed auto also have "\ = ennreal (1 - (\n\k. (a^n * exp (-a)) / fact n))" by (auto simp: power_0_left if_distrib[where f="\x. x / a" for a] sum.If_cases) also have "\ = ennreal ((1 - (\n\k. (a^n * exp (-a)) / fact n)) * fact k) * ennreal (inverse (fact k))" by (subst ennreal_mult''[symmetric]) (auto intro!: arg_cong[where f=ennreal]) finally show ?thesis by (auto simp add: mult_right_ennreal_cancel le_less) qed lemma nn_intergal_power_times_exp_Ici: shows "(\\<^sup>+x. ennreal (x^k * exp (-x)) * indicator {0 ..} x \lborel) = real_of_nat (fact k)" proof (rule LIMSEQ_unique) let ?X = "\n. \\<^sup>+ x. ennreal (x^k * exp (-x)) * indicator {0 .. real n} x \lborel" show "?X \ (\\<^sup>+x. ennreal (x^k * exp (-x)) * indicator {0 ..} x \lborel)" apply (intro nn_integral_LIMSEQ) apply (auto simp: incseq_def le_fun_def eventually_sequentially split: split_indicator intro!: tendsto_eventually) apply (metis nat_ceiling_le_eq) done have "((\x::real. (1 - (\n\k. (x ^ n / exp x) / (fact n))) * fact k) \ (1 - (\n\k. 0 / (fact n))) * fact k) at_top" by (intro tendsto_intros tendsto_power_div_exp_0) simp then show "?X \ real_of_nat (fact k)" by (subst nn_intergal_power_times_exp_Icc) (auto simp: exp_minus field_simps intro!: filterlim_compose[OF _ filterlim_real_sequentially]) qed definition erlang_density :: "nat \ real \ real \ real" where "erlang_density k l x = (if x < 0 then 0 else (l^(Suc k) * x^k * exp (- l * x)) / fact k)" definition erlang_CDF :: "nat \ real \ real \ real" where "erlang_CDF k l x = (if x < 0 then 0 else 1 - (\n\k. ((l * x)^n * exp (- l * x) / fact n)))" lemma erlang_density_nonneg[simp]: "0 \ l \ 0 \ erlang_density k l x" by (simp add: erlang_density_def) lemma borel_measurable_erlang_density[measurable]: "erlang_density k l \ borel_measurable borel" by (auto simp add: erlang_density_def[abs_def]) lemma erlang_CDF_transform: "0 < l \ erlang_CDF k l a = erlang_CDF k 1 (l * a)" by (auto simp add: erlang_CDF_def mult_less_0_iff) lemma erlang_CDF_nonneg[simp]: assumes "0 < l" shows "0 \ erlang_CDF k l x" unfolding erlang_CDF_def proof (clarsimp simp: not_less) assume "0 \ x" have "(\n\k. (l * x) ^ n * exp (- (l * x)) / fact n) = exp (- (l * x)) * (\n\k. (l * x) ^ n / fact n)" unfolding sum_distrib_left by (intro sum.cong) (auto simp: field_simps) also have "\ = (\n\k. (l * x) ^ n / fact n) / exp (l * x)" by (simp add: exp_minus field_simps) also have "\ \ 1" proof (subst divide_le_eq_1_pos) show "(\n\k. (l * x) ^ n / fact n) \ exp (l * x)" using \0 < l\ \0 \ x\ summable_exp_generic[of "l * x"] by (auto simp: exp_def divide_inverse ac_simps intro!: sum_le_suminf) qed simp finally show "(\n\k. (l * x) ^ n * exp (- (l * x)) / fact n) \ 1" . qed lemma nn_integral_erlang_density: assumes [arith]: "0 < l" shows "(\\<^sup>+ x. ennreal (erlang_density k l x) * indicator {.. a} x \lborel) = erlang_CDF k l a" proof (cases "0 \ a") case [arith]: True have eq: "\x. indicator {0..a} (x / l) = indicator {0..a*l} x" by (simp add: field_simps split: split_indicator) have "(\\<^sup>+x. ennreal (erlang_density k l x) * indicator {.. a} x \lborel) = (\\<^sup>+x. (l/fact k) * (ennreal ((l*x)^k * exp (- (l*x))) * indicator {0 .. a} x) \lborel)" by (intro nn_integral_cong) (auto simp: erlang_density_def power_mult_distrib ennreal_mult[symmetric] split: split_indicator) also have "\ = (l/fact k) * (\\<^sup>+x. ennreal ((l*x)^k * exp (- (l*x))) * indicator {0 .. a} x \lborel)" by (intro nn_integral_cmult) auto also have "\ = ennreal (l/fact k) * ((1/l) * (\\<^sup>+x. ennreal (x^k * exp (- x)) * indicator {0 .. l * a} x \lborel))" by (subst nn_integral_real_affine[where c="1 / l" and t=0]) (auto simp: field_simps eq) also have "\ = (1 - (\n\k. ((l * a)^n * exp (-(l * a))) / fact n))" by (subst nn_intergal_power_times_exp_Icc) (auto simp: ennreal_mult'[symmetric]) also have "\ = erlang_CDF k l a" by (auto simp: erlang_CDF_def) finally show ?thesis . next case False then have "(\\<^sup>+ x. ennreal (erlang_density k l x) * indicator {.. a} x \lborel) = (\\<^sup>+x. 0 \(lborel::real measure))" by (intro nn_integral_cong) (auto simp: erlang_density_def) with False show ?thesis by (simp add: erlang_CDF_def) qed lemma emeasure_erlang_density: "0 < l \ emeasure (density lborel (erlang_density k l)) {.. a} = erlang_CDF k l a" by (simp add: emeasure_density nn_integral_erlang_density) lemma nn_integral_erlang_ith_moment: fixes k i :: nat and l :: real assumes [arith]: "0 < l" shows "(\\<^sup>+ x. ennreal (erlang_density k l x * x ^ i) \lborel) = fact (k + i) / (fact k * l ^ i)" proof - have eq: "\x. indicator {0..} (x / l) = indicator {0..} x" by (simp add: field_simps split: split_indicator) have "(\\<^sup>+ x. ennreal (erlang_density k l x * x^i) \lborel) = (\\<^sup>+x. (l/(fact k * l^i)) * (ennreal ((l*x)^(k+i) * exp (- (l*x))) * indicator {0 ..} x) \lborel)" by (intro nn_integral_cong) (auto simp: erlang_density_def power_mult_distrib power_add ennreal_mult'[symmetric] split: split_indicator) also have "\ = (l/(fact k * l^i)) * (\\<^sup>+x. ennreal ((l*x)^(k+i) * exp (- (l*x))) * indicator {0 ..} x \lborel)" by (intro nn_integral_cmult) auto also have "\ = ennreal (l/(fact k * l^i)) * ((1/l) * (\\<^sup>+x. ennreal (x^(k+i) * exp (- x)) * indicator {0 ..} x \lborel))" by (subst nn_integral_real_affine[where c="1 / l" and t=0]) (auto simp: field_simps eq) also have "\ = fact (k + i) / (fact k * l ^ i)" by (subst nn_intergal_power_times_exp_Ici) (auto simp: ennreal_mult'[symmetric]) finally show ?thesis . qed lemma prob_space_erlang_density: assumes l[arith]: "0 < l" shows "prob_space (density lborel (erlang_density k l))" (is "prob_space ?D") proof show "emeasure ?D (space ?D) = 1" using nn_integral_erlang_ith_moment[OF l, where k=k and i=0] by (simp add: emeasure_density) qed lemma (in prob_space) erlang_distributed_le: assumes D: "distributed M lborel X (erlang_density k l)" assumes [simp, arith]: "0 < l" "0 \ a" shows "\

(x in M. X x \ a) = erlang_CDF k l a" proof - have "emeasure M {x \ space M. X x \ a } = emeasure (distr M lborel X) {.. a}" using distributed_measurable[OF D] by (subst emeasure_distr) (auto intro!: arg_cong2[where f=emeasure]) also have "\ = emeasure (density lborel (erlang_density k l)) {.. a}" unfolding distributed_distr_eq_density[OF D] .. also have "\ = erlang_CDF k l a" by (auto intro!: emeasure_erlang_density) finally show ?thesis by (auto simp: emeasure_eq_measure measure_nonneg) qed lemma (in prob_space) erlang_distributed_gt: assumes D[simp]: "distributed M lborel X (erlang_density k l)" assumes [arith]: "0 < l" "0 \ a" shows "\

(x in M. a < X x ) = 1 - (erlang_CDF k l a)" proof - have " 1 - (erlang_CDF k l a) = 1 - \

(x in M. X x \ a)" by (subst erlang_distributed_le) auto also have "\ = prob (space M - {x \ space M. X x \ a })" using distributed_measurable[OF D] by (auto simp: prob_compl) also have "\ = \

(x in M. a < X x )" by (auto intro!: arg_cong[where f=prob] simp: not_le) finally show ?thesis by simp qed lemma erlang_CDF_at0: "erlang_CDF k l 0 = 0" by (induction k) (auto simp: erlang_CDF_def) lemma erlang_distributedI: assumes X[measurable]: "X \ borel_measurable M" and [arith]: "0 < l" and X_distr: "\a. 0 \ a \ emeasure M {x\space M. X x \ a} = erlang_CDF k l a" shows "distributed M lborel X (erlang_density k l)" proof (rule distributedI_borel_atMost) fix a :: real { assume "a \ 0" with X have "emeasure M {x\space M. X x \ a} \ emeasure M {x\space M. X x \ 0}" by (intro emeasure_mono) auto also have "... = 0" by (auto intro!: erlang_CDF_at0 simp: X_distr[of 0]) finally have "emeasure M {x\space M. X x \ a} \ 0" by simp then have "emeasure M {x\space M. X x \ a} = 0" by simp } note eq_0 = this show "(\\<^sup>+ x. erlang_density k l x * indicator {..a} x \lborel) = ennreal (erlang_CDF k l a)" using nn_integral_erlang_density[of l k a] by (simp add: ennreal_indicator ennreal_mult) show "emeasure M {x\space M. X x \ a} = ennreal (erlang_CDF k l a)" using X_distr[of a] eq_0 by (auto simp: one_ennreal_def erlang_CDF_def) qed simp_all lemma (in prob_space) erlang_distributed_iff: assumes [arith]: "0 (X \ borel_measurable M \ 0 < l \ (\a\0. \

(x in M. X x \ a) = erlang_CDF k l a ))" using distributed_measurable[of M lborel X "erlang_density k l"] emeasure_erlang_density[of l] erlang_distributed_le[of X k l] by (auto intro!: erlang_distributedI simp: one_ennreal_def emeasure_eq_measure) lemma (in prob_space) erlang_distributed_mult_const: assumes erlX: "distributed M lborel X (erlang_density k l)" assumes a_pos[arith]: "0 < \" "0 < l" shows "distributed M lborel (\x. \ * X x) (erlang_density k (l / \))" proof (subst erlang_distributed_iff, safe) have [measurable]: "random_variable borel X" and [arith]: "0 < l " and [simp]: "\a. 0 \ a \ prob {x \ space M. X x \ a} = erlang_CDF k l a" by(insert erlX, auto simp: erlang_distributed_iff) show "random_variable borel (\x. \ * X x)" "0 < l / \" "0 < l / \" by (auto simp:field_simps) fix a:: real assume [arith]: "0 \ a" obtain b:: real where [simp, arith]: "b = a/ \" by blast have [arith]: "0 \ b" by (auto simp: divide_nonneg_pos) have "prob {x \ space M. \ * X x \ a} = prob {x \ space M. X x \ b}" by (rule arg_cong[where f= prob]) (auto simp:field_simps) moreover have "prob {x \ space M. X x \ b} = erlang_CDF k l b" by auto moreover have "erlang_CDF k (l / \) a = erlang_CDF k l b" unfolding erlang_CDF_def by auto ultimately show "prob {x \ space M. \ * X x \ a} = erlang_CDF k (l / \) a" by fastforce qed lemma (in prob_space) has_bochner_integral_erlang_ith_moment: fixes k i :: nat and l :: real assumes [arith]: "0 < l" and D: "distributed M lborel X (erlang_density k l)" shows "has_bochner_integral M (\x. X x ^ i) (fact (k + i) / (fact k * l ^ i))" proof (rule has_bochner_integral_nn_integral) show "AE x in M. 0 \ X x ^ i" by (subst distributed_AE2[OF D]) (auto simp: erlang_density_def) show "(\\<^sup>+ x. ennreal (X x ^ i) \M) = ennreal (fact (k + i) / (fact k * l ^ i))" using nn_integral_erlang_ith_moment[of l k i] by (subst distributed_nn_integral[symmetric, OF D]) (auto simp: ennreal_mult') qed (insert distributed_measurable[OF D], auto) lemma (in prob_space) erlang_ith_moment_integrable: "0 < l \ distributed M lborel X (erlang_density k l) \ integrable M (\x. X x ^ i)" by rule (rule has_bochner_integral_erlang_ith_moment) lemma (in prob_space) erlang_ith_moment: "0 < l \ distributed M lborel X (erlang_density k l) \ expectation (\x. X x ^ i) = fact (k + i) / (fact k * l ^ i)" by (rule has_bochner_integral_integral_eq) (rule has_bochner_integral_erlang_ith_moment) lemma (in prob_space) erlang_distributed_variance: assumes [arith]: "0 < l" and "distributed M lborel X (erlang_density k l)" shows "variance X = (k + 1) / l\<^sup>2" proof (subst variance_eq) show "integrable M X" "integrable M (\x. (X x)\<^sup>2)" using erlang_ith_moment_integrable[OF assms, of 1] erlang_ith_moment_integrable[OF assms, of 2] by auto show "expectation (\x. (X x)\<^sup>2) - (expectation X)\<^sup>2 = real (k + 1) / l\<^sup>2" using erlang_ith_moment[OF assms, of 1] erlang_ith_moment[OF assms, of 2] by simp (auto simp: power2_eq_square field_simps of_nat_Suc) qed subsection \Exponential distribution\ abbreviation exponential_density :: "real \ real \ real" where "exponential_density \ erlang_density 0" lemma exponential_density_def: "exponential_density l x = (if x < 0 then 0 else l * exp (- x * l))" by (simp add: fun_eq_iff erlang_density_def) lemma erlang_CDF_0: "erlang_CDF 0 l a = (if 0 \ a then 1 - exp (- l * a) else 0)" by (simp add: erlang_CDF_def) lemma prob_space_exponential_density: "0 < l \ prob_space (density lborel (exponential_density l))" by (rule prob_space_erlang_density) lemma (in prob_space) exponential_distributedD_le: assumes D: "distributed M lborel X (exponential_density l)" and a: "0 \ a" and l: "0 < l" shows "\

(x in M. X x \ a) = 1 - exp (- a * l)" using erlang_distributed_le[OF D l a] a by (simp add: erlang_CDF_def) lemma (in prob_space) exponential_distributedD_gt: assumes D: "distributed M lborel X (exponential_density l)" and a: "0 \ a" and l: "0 < l" shows "\

(x in M. a < X x ) = exp (- a * l)" using erlang_distributed_gt[OF D l a] a by (simp add: erlang_CDF_def) lemma (in prob_space) exponential_distributed_memoryless: assumes D: "distributed M lborel X (exponential_density l)" and a: "0 \ a" and l: "0 < l" and t: "0 \ t" shows "\

(x in M. a + t < X x \ a < X x) = \

(x in M. t < X x)" proof - have "\

(x in M. a + t < X x \ a < X x) = \

(x in M. a + t < X x) / \

(x in M. a < X x)" using \0 \ t\ by (auto simp: cond_prob_def intro!: arg_cong[where f=prob] arg_cong2[where f="(/)"]) also have "\ = exp (- (a + t) * l) / exp (- a * l)" using a t by (simp add: exponential_distributedD_gt[OF D _ l]) also have "\ = exp (- t * l)" using l by (auto simp: field_simps exp_add[symmetric]) finally show ?thesis using t by (simp add: exponential_distributedD_gt[OF D _ l]) qed lemma exponential_distributedI: assumes X[measurable]: "X \ borel_measurable M" and [arith]: "0 < l" and X_distr: "\a. 0 \ a \ emeasure M {x\space M. X x \ a} = 1 - exp (- a * l)" shows "distributed M lborel X (exponential_density l)" proof (rule erlang_distributedI) fix a :: real assume "0 \ a" then show "emeasure M {x \ space M. X x \ a} = ennreal (erlang_CDF 0 l a)" using X_distr[of a] by (simp add: erlang_CDF_def ennreal_minus ennreal_1[symmetric] del: ennreal_1) qed fact+ lemma (in prob_space) exponential_distributed_iff: assumes "0 < l" shows "distributed M lborel X (exponential_density l) \ (X \ borel_measurable M \ (\a\0. \

(x in M. X x \ a) = 1 - exp (- a * l)))" using assms erlang_distributed_iff[of l X 0] by (auto simp: erlang_CDF_0) lemma (in prob_space) exponential_distributed_expectation: "0 < l \ distributed M lborel X (exponential_density l) \ expectation X = 1 / l" using erlang_ith_moment[of l X 0 1] by simp lemma exponential_density_nonneg: "0 < l \ 0 \ exponential_density l x" by (auto simp: exponential_density_def) lemma (in prob_space) exponential_distributed_min: assumes "0 < l" "0 < u" assumes expX: "distributed M lborel X (exponential_density l)" assumes expY: "distributed M lborel Y (exponential_density u)" assumes ind: "indep_var borel X borel Y" shows "distributed M lborel (\x. min (X x) (Y x)) (exponential_density (l + u))" proof (subst exponential_distributed_iff, safe) have randX: "random_variable borel X" using expX \0 < l\ by (simp add: exponential_distributed_iff) moreover have randY: "random_variable borel Y" using expY \0 < u\ by (simp add: exponential_distributed_iff) ultimately show "random_variable borel (\x. min (X x) (Y x))" by auto show " 0 < l + u" using \0 < l\ \0 < u\ by auto fix a::real assume a[arith]: "0 \ a" have gt1[simp]: "\

(x in M. a < X x ) = exp (- a * l)" by (rule exponential_distributedD_gt[OF expX a]) fact have gt2[simp]: "\

(x in M. a < Y x ) = exp (- a * u)" by (rule exponential_distributedD_gt[OF expY a]) fact have "\

(x in M. a < (min (X x) (Y x)) ) = \

(x in M. a < (X x) \ a < (Y x))" by (auto intro!:arg_cong[where f=prob]) also have " ... = \

(x in M. a < (X x)) * \

(x in M. a< (Y x) )" using prob_indep_random_variable[OF ind, of "{a <..}" "{a <..}"] by simp also have " ... = exp (- a * (l + u))" by (auto simp:field_simps mult_exp_exp) finally have indep_prob: "\

(x in M. a < (min (X x) (Y x)) ) = exp (- a * (l + u))" . have "{x \ space M. (min (X x) (Y x)) \a } = (space M - {x \ space M. a<(min (X x) (Y x)) })" by auto then have "1 - prob {x \ space M. a < (min (X x) (Y x))} = prob {x \ space M. (min (X x) (Y x)) \ a}" using randX randY by (auto simp: prob_compl) then show "prob {x \ space M. (min (X x) (Y x)) \ a} = 1 - exp (- a * (l + u))" using indep_prob by auto qed lemma (in prob_space) exponential_distributed_Min: assumes finI: "finite I" assumes A: "I \ {}" assumes l: "\i. i \ I \ 0 < l i" assumes expX: "\i. i \ I \ distributed M lborel (X i) (exponential_density (l i))" assumes ind: "indep_vars (\i. borel) X I" shows "distributed M lborel (\x. Min ((\i. X i x)`I)) (exponential_density (\i\I. l i))" using assms proof (induct rule: finite_ne_induct) case (singleton i) then show ?case by simp next case (insert i I) then have "distributed M lborel (\x. min (X i x) (Min ((\i. X i x)`I))) (exponential_density (l i + (\i\I. l i)))" by (intro exponential_distributed_min indep_vars_Min insert) (auto intro: indep_vars_subset sum_pos) then show ?case using insert by simp qed lemma (in prob_space) exponential_distributed_variance: "0 < l \ distributed M lborel X (exponential_density l) \ variance X = 1 / l\<^sup>2" using erlang_distributed_variance[of l X 0] by simp lemma nn_integral_zero': "AE x in M. f x = 0 \ (\\<^sup>+x. f x \M) = 0" by (simp cong: nn_integral_cong_AE) lemma convolution_erlang_density: fixes k\<^sub>1 k\<^sub>2 :: nat assumes [simp, arith]: "0 < l" shows "(\x. \\<^sup>+y. ennreal (erlang_density k\<^sub>1 l (x - y)) * ennreal (erlang_density k\<^sub>2 l y) \lborel) = (erlang_density (Suc k\<^sub>1 + Suc k\<^sub>2 - 1) l)" (is "?LHS = ?RHS") proof fix x :: real have "x \ 0 \ 0 < x" by arith then show "?LHS x = ?RHS x" proof assume "x \ 0" then show ?thesis apply (subst nn_integral_zero') apply (rule AE_I[where N="{0}"]) apply (auto simp add: erlang_density_def not_less) done next note zero_le_mult_iff[simp] zero_le_divide_iff[simp] have I_eq1: "integral\<^sup>N lborel (erlang_density (Suc k\<^sub>1 + Suc k\<^sub>2 - 1) l) = 1" using nn_integral_erlang_ith_moment[of l "Suc k\<^sub>1 + Suc k\<^sub>2 - 1" 0] by (simp del: fact_Suc) have 1: "(\\<^sup>+ x. ennreal (erlang_density (Suc k\<^sub>1 + Suc k\<^sub>2 - 1) l x * indicator {0<..} x) \lborel) = 1" apply (subst I_eq1[symmetric]) unfolding erlang_density_def by (auto intro!: nn_integral_cong split:split_indicator) have "prob_space (density lborel ?LHS)" by (intro prob_space_convolution_density) (auto intro!: prob_space_erlang_density erlang_density_nonneg) then have 2: "integral\<^sup>N lborel ?LHS = 1" by (auto dest!: prob_space.emeasure_space_1 simp: emeasure_density) let ?I = "(integral\<^sup>N lborel (\y. ennreal ((1 - y)^ k\<^sub>1 * y^k\<^sub>2 * indicator {0..1} y)))" let ?C = "(fact (Suc (k\<^sub>1 + k\<^sub>2))) / ((fact k\<^sub>1) * (fact k\<^sub>2))" let ?s = "Suc k\<^sub>1 + Suc k\<^sub>2 - 1" let ?L = "(\x. \\<^sup>+y. ennreal (erlang_density k\<^sub>1 l (x- y) * erlang_density k\<^sub>2 l y * indicator {0..x} y) \lborel)" { fix x :: real assume [arith]: "0 < x" have *: "\x y n. (x - y * x::real)^n = x^n * (1 - y)^n" unfolding power_mult_distrib[symmetric] by (simp add: field_simps) have "?LHS x = ?L x" unfolding erlang_density_def by (auto intro!: nn_integral_cong simp: ennreal_mult split:split_indicator) also have "... = (\x. ennreal ?C * ?I * erlang_density ?s l x) x" apply (subst nn_integral_real_affine[where c=x and t = 0]) apply (simp_all add: nn_integral_cmult[symmetric] nn_integral_multc[symmetric] del: fact_Suc) apply (intro nn_integral_cong) apply (auto simp add: erlang_density_def mult_less_0_iff exp_minus field_simps exp_diff power_add * ennreal_mult[symmetric] simp del: fact_Suc split: split_indicator) done finally have "(\\<^sup>+y. ennreal (erlang_density k\<^sub>1 l (x - y) * erlang_density k\<^sub>2 l y) \lborel) = (\x. ennreal ?C * ?I * erlang_density ?s l x) x" by (simp add: ennreal_mult) } note * = this assume [arith]: "0 < x" have 3: "1 = integral\<^sup>N lborel (\xa. ?LHS xa * indicator {0<..} xa)" by (subst 2[symmetric]) (auto intro!: nn_integral_cong_AE AE_I[where N="{0}"] simp: erlang_density_def nn_integral_multc[symmetric] indicator_def split: if_split_asm) also have "... = integral\<^sup>N lborel (\x. (ennreal (?C) * ?I) * ((erlang_density ?s l x) * indicator {0<..} x))" by (auto intro!: nn_integral_cong simp: ennreal_mult[symmetric] * split: split_indicator) also have "... = ennreal (?C) * ?I" using 1 by (auto simp: nn_integral_cmult) finally have " ennreal (?C) * ?I = 1" by presburger then show ?thesis using * by (simp add: ennreal_mult) qed qed lemma (in prob_space) sum_indep_erlang: assumes indep: "indep_var borel X borel Y" assumes [simp, arith]: "0 < l" assumes erlX: "distributed M lborel X (erlang_density k\<^sub>1 l)" assumes erlY: "distributed M lborel Y (erlang_density k\<^sub>2 l)" shows "distributed M lborel (\x. X x + Y x) (erlang_density (Suc k\<^sub>1 + Suc k\<^sub>2 - 1) l)" using assms apply (subst convolution_erlang_density[symmetric, OF \0]) apply (intro distributed_convolution) apply auto done lemma (in prob_space) erlang_distributed_sum: assumes finI : "finite I" assumes A: "I \ {}" assumes [simp, arith]: "0 < l" assumes expX: "\i. i \ I \ distributed M lborel (X i) (erlang_density (k i) l)" assumes ind: "indep_vars (\i. borel) X I" shows "distributed M lborel (\x. \i\I. X i x) (erlang_density ((\i\I. Suc (k i)) - 1) l)" using assms proof (induct rule: finite_ne_induct) case (singleton i) then show ?case by auto next case (insert i I) then have "distributed M lborel (\x. (X i x) + (\i\ I. X i x)) (erlang_density (Suc (k i) + Suc ((\i\I. Suc (k i)) - 1) - 1) l)" by(intro sum_indep_erlang indep_vars_sum) (auto intro!: indep_vars_subset) also have "(\x. (X i x) + (\i\ I. X i x)) = (\x. \i\insert i I. X i x)" using insert by auto also have "Suc(k i) + Suc ((\i\I. Suc (k i)) - 1) - 1 = (\i\insert i I. Suc (k i)) - 1" using insert by (auto intro!: Suc_pred simp: ac_simps) finally show ?case by fast qed lemma (in prob_space) exponential_distributed_sum: assumes finI: "finite I" assumes A: "I \ {}" assumes l: "0 < l" assumes expX: "\i. i \ I \ distributed M lborel (X i) (exponential_density l)" assumes ind: "indep_vars (\i. borel) X I" shows "distributed M lborel (\x. \i\I. X i x) (erlang_density ((card I) - 1) l)" using erlang_distributed_sum[OF assms] by simp lemma (in information_space) entropy_exponential: assumes l[simp, arith]: "0 < l" assumes D: "distributed M lborel X (exponential_density l)" shows "entropy b lborel X = log b (exp 1 / l)" proof - have [simp]: "integrable lborel (exponential_density l)" using distributed_integrable[OF D, of "\_. 1"] by simp have [simp]: "integral\<^sup>L lborel (exponential_density l) = 1" using distributed_integral[OF D, of "\_. 1"] by (simp add: prob_space) have [simp]: "integrable lborel (\x. exponential_density l x * x)" using erlang_ith_moment_integrable[OF l D, of 1] distributed_integrable[OF D, of "\x. x"] by simp have [simp]: "integral\<^sup>L lborel (\x. exponential_density l x * x) = 1 / l" using erlang_ith_moment[OF l D, of 1] distributed_integral[OF D, of "\x. x"] by simp have "entropy b lborel X = - (\ x. exponential_density l x * log b (exponential_density l x) \lborel)" using D by (rule entropy_distr) simp also have "(\ x. exponential_density l x * log b (exponential_density l x) \lborel) = (\ x. (ln l * exponential_density l x - l * (exponential_density l x * x)) / ln b \lborel)" by (intro Bochner_Integration.integral_cong) (auto simp: log_def ln_mult exponential_density_def field_simps) also have "\ = (ln l - 1) / ln b" by simp finally show ?thesis by (simp add: log_def ln_div) (simp add: field_split_simps) qed subsection \Uniform distribution\ lemma uniform_distrI: assumes X: "X \ measurable M M'" and A: "A \ sets M'" "emeasure M' A \ \" "emeasure M' A \ 0" assumes distr: "\B. B \ sets M' \ emeasure M (X -` B \ space M) = emeasure M' (A \ B) / emeasure M' A" shows "distr M M' X = uniform_measure M' A" unfolding uniform_measure_def proof (intro measure_eqI) let ?f = "\x. indicator A x / emeasure M' A" fix B assume B: "B \ sets (distr M M' X)" with X have "emeasure M (X -` B \ space M) = emeasure M' (A \ B) / emeasure M' A" by (simp add: distr[of B] measurable_sets) also have "\ = (1 / emeasure M' A) * emeasure M' (A \ B)" by (simp add: divide_ennreal_def ac_simps) also have "\ = (\\<^sup>+ x. (1 / emeasure M' A) * indicator (A \ B) x \M')" using A B by (intro nn_integral_cmult_indicator[symmetric]) (auto intro!: ) also have "\ = (\\<^sup>+ x. ?f x * indicator B x \M')" by (rule nn_integral_cong) (auto split: split_indicator) finally show "emeasure (distr M M' X) B = emeasure (density M' ?f) B" using A B X by (auto simp add: emeasure_distr emeasure_density) qed simp lemma uniform_distrI_borel: fixes A :: "real set" assumes X[measurable]: "X \ borel_measurable M" and A: "emeasure lborel A = ennreal r" "0 < r" and [measurable]: "A \ sets borel" assumes distr: "\a. emeasure M {x\space M. X x \ a} = emeasure lborel (A \ {.. a}) / r" shows "distributed M lborel X (\x. indicator A x / measure lborel A)" proof (rule distributedI_borel_atMost) let ?f = "\x. 1 / r * indicator A x" fix a have "emeasure lborel (A \ {..a}) \ emeasure lborel A" using A by (intro emeasure_mono) auto also have "\ < \" using A by simp finally have fin: "emeasure lborel (A \ {..a}) \ top" by simp from emeasure_eq_ennreal_measure[OF this] have fin_eq: "emeasure lborel (A \ {..a}) / r = ennreal (measure lborel (A \ {..a}) / r)" using A by (simp add: divide_ennreal measure_nonneg) then show "emeasure M {x\space M. X x \ a} = ennreal (measure lborel (A \ {..a}) / r)" using distr by simp have "(\\<^sup>+ x. ennreal (indicator A x / measure lborel A * indicator {..a} x) \lborel) = (\\<^sup>+ x. ennreal (1 / measure lborel A) * indicator (A \ {..a}) x \lborel)" by (auto intro!: nn_integral_cong split: split_indicator) also have "\ = ennreal (1 / measure lborel A) * emeasure lborel (A \ {..a})" using \A \ sets borel\ by (intro nn_integral_cmult_indicator) (auto simp: measure_nonneg) also have "\ = ennreal (measure lborel (A \ {..a}) / r)" unfolding emeasure_eq_ennreal_measure[OF fin] using A by (simp add: measure_def ennreal_mult'[symmetric]) finally show "(\\<^sup>+ x. ennreal (indicator A x / measure lborel A * indicator {..a} x) \lborel) = ennreal (measure lborel (A \ {..a}) / r)" . qed (auto simp: measure_nonneg) lemma (in prob_space) uniform_distrI_borel_atLeastAtMost: fixes a b :: real assumes X: "X \ borel_measurable M" and "a < b" assumes distr: "\t. a \ t \ t \ b \ \

(x in M. X x \ t) = (t - a) / (b - a)" shows "distributed M lborel X (\x. indicator {a..b} x / measure lborel {a..b})" proof (rule uniform_distrI_borel) fix t have "t < a \ (a \ t \ t \ b) \ b < t" by auto then show "emeasure M {x\space M. X x \ t} = emeasure lborel ({a .. b} \ {..t}) / (b - a)" proof (elim disjE conjE) assume "t < a" then have "emeasure M {x\space M. X x \ t} \ emeasure M {x\space M. X x \ a}" using X by (auto intro!: emeasure_mono measurable_sets) also have "\ = 0" using distr[of a] \a < b\ by (simp add: emeasure_eq_measure) finally have "emeasure M {x\space M. X x \ t} = 0" by (simp add: antisym measure_nonneg) with \t < a\ show ?thesis by simp next assume bnds: "a \ t" "t \ b" have "{a..b} \ {..t} = {a..t}" using bnds by auto then show ?thesis using \a \ t\ \a < b\ using distr[OF bnds] by (simp add: emeasure_eq_measure divide_ennreal) next assume "b < t" have "1 = emeasure M {x\space M. X x \ b}" using distr[of b] \a < b\ by (simp add: one_ennreal_def emeasure_eq_measure) also have "\ \ emeasure M {x\space M. X x \ t}" using X \b < t\ by (auto intro!: emeasure_mono measurable_sets) finally have "emeasure M {x\space M. X x \ t} = 1" by (simp add: antisym emeasure_eq_measure) with \b < t\ \a < b\ show ?thesis by (simp add: measure_def divide_ennreal) qed qed (insert X \a < b\, auto) lemma (in prob_space) uniform_distributed_measure: fixes a b :: real assumes D: "distributed M lborel X (\x. indicator {a .. b} x / measure lborel {a .. b})" assumes t: "a \ t" "t \ b" shows "\

(x in M. X x \ t) = (t - a) / (b - a)" proof - have "emeasure M {x \ space M. X x \ t} = emeasure (distr M lborel X) {.. t}" using distributed_measurable[OF D] by (subst emeasure_distr) (auto intro!: arg_cong2[where f=emeasure]) also have "\ = (\\<^sup>+x. ennreal (1 / (b - a)) * indicator {a .. t} x \lborel)" using distributed_borel_measurable[OF D] \a \ t\ \t \ b\ unfolding distributed_distr_eq_density[OF D] by (subst emeasure_density) (auto intro!: nn_integral_cong simp: measure_def split: split_indicator) also have "\ = ennreal (1 / (b - a)) * (t - a)" using \a \ t\ \t \ b\ by (subst nn_integral_cmult_indicator) auto finally show ?thesis using t by (simp add: emeasure_eq_measure ennreal_mult''[symmetric] measure_nonneg) qed lemma (in prob_space) uniform_distributed_bounds: fixes a b :: real assumes D: "distributed M lborel X (\x. indicator {a .. b} x / measure lborel {a .. b})" shows "a < b" proof (rule ccontr) assume "\ a < b" then have "{a .. b} = {} \ {a .. b} = {a .. a}" by simp with uniform_distributed_params[OF D] show False by (auto simp: measure_def) qed lemma (in prob_space) uniform_distributed_iff: fixes a b :: real shows "distributed M lborel X (\x. indicator {a..b} x / measure lborel {a..b}) \ (X \ borel_measurable M \ a < b \ (\t\{a .. b}. \

(x in M. X x \ t)= (t - a) / (b - a)))" using uniform_distributed_bounds[of X a b] uniform_distributed_measure[of X a b] distributed_measurable[of M lborel X] by (auto intro!: uniform_distrI_borel_atLeastAtMost simp del: content_real_if) lemma (in prob_space) uniform_distributed_expectation: fixes a b :: real assumes D: "distributed M lborel X (\x. indicator {a .. b} x / measure lborel {a .. b})" shows "expectation X = (a + b) / 2" proof (subst distributed_integral[OF D, of "\x. x", symmetric]) have "a < b" using uniform_distributed_bounds[OF D] . have "(\ x. indicator {a .. b} x / measure lborel {a .. b} * x \lborel) = (\ x. (x / measure lborel {a .. b}) * indicator {a .. b} x \lborel)" by (intro Bochner_Integration.integral_cong) auto also have "(\ x. (x / measure lborel {a .. b}) * indicator {a .. b} x \lborel) = (a + b) / 2" proof (subst integral_FTC_Icc_real) fix x show "DERIV (\x. x\<^sup>2 / (2 * measure lborel {a..b})) x :> x / measure lborel {a..b}" using uniform_distributed_params[OF D] by (auto intro!: derivative_eq_intros simp del: content_real_if) show "isCont (\x. x / Sigma_Algebra.measure lborel {a..b}) x" using uniform_distributed_params[OF D] by (auto intro!: isCont_divide) have *: "b\<^sup>2 / (2 * measure lborel {a..b}) - a\<^sup>2 / (2 * measure lborel {a..b}) = (b*b - a * a) / (2 * (b - a))" using \a < b\ by (auto simp: measure_def power2_eq_square diff_divide_distrib[symmetric]) show "b\<^sup>2 / (2 * measure lborel {a..b}) - a\<^sup>2 / (2 * measure lborel {a..b}) = (a + b) / 2" using \a < b\ unfolding * square_diff_square_factored by (auto simp: field_simps) qed (insert \a < b\, simp) finally show "(\ x. indicator {a .. b} x / measure lborel {a .. b} * x \lborel) = (a + b) / 2" . qed (auto simp: measure_nonneg) lemma (in prob_space) uniform_distributed_variance: fixes a b :: real assumes D: "distributed M lborel X (\x. indicator {a .. b} x / measure lborel {a .. b})" shows "variance X = (b - a)\<^sup>2 / 12" proof (subst distributed_variance) have [arith]: "a < b" using uniform_distributed_bounds[OF D] . let ?\ = "expectation X" let ?D = "\x. indicator {a..b} (x + ?\) / measure lborel {a..b}" have "(\x. x\<^sup>2 * (?D x) \lborel) = (\x. x\<^sup>2 * (indicator {a - ?\ .. b - ?\} x) / measure lborel {a .. b} \lborel)" by (intro Bochner_Integration.integral_cong) (auto split: split_indicator) also have "\ = (b - a)\<^sup>2 / 12" by (simp add: integral_power uniform_distributed_expectation[OF D]) (simp add: eval_nat_numeral field_simps ) finally show "(\x. x\<^sup>2 * ?D x \lborel) = (b - a)\<^sup>2 / 12" . qed (auto intro: D simp del: content_real_if) subsection \Normal distribution\ definition normal_density :: "real \ real \ real \ real" where "normal_density \ \ x = 1 / sqrt (2 * pi * \\<^sup>2) * exp (-(x - \)\<^sup>2/ (2 * \\<^sup>2))" abbreviation std_normal_density :: "real \ real" where "std_normal_density \ normal_density 0 1" lemma std_normal_density_def: "std_normal_density x = (1 / sqrt (2 * pi)) * exp (- x\<^sup>2 / 2)" unfolding normal_density_def by simp lemma normal_density_nonneg[simp]: "0 \ normal_density \ \ x" by (auto simp: normal_density_def) lemma normal_density_pos: "0 < \ \ 0 < normal_density \ \ x" by (auto simp: normal_density_def) lemma borel_measurable_normal_density[measurable]: "normal_density \ \ \ borel_measurable borel" by (auto simp: normal_density_def[abs_def]) lemma gaussian_moment_0: "has_bochner_integral lborel (\x. indicator {0..} x *\<^sub>R exp (- x\<^sup>2)) (sqrt pi / 2)" proof - let ?pI = "\f. (\\<^sup>+s. f (s::real) * indicator {0..} s \lborel)" let ?gauss = "\x. exp (- x\<^sup>2)" let ?I = "indicator {0<..} :: real \ real" let ?ff= "\x s. x * exp (- x\<^sup>2 * (1 + s\<^sup>2)) :: real" have *: "?pI ?gauss = (\\<^sup>+x. ?gauss x * ?I x \lborel)" by (intro nn_integral_cong_AE AE_I[where N="{0}"]) (auto split: split_indicator) have "?pI ?gauss * ?pI ?gauss = (\\<^sup>+x. \\<^sup>+s. ?gauss x * ?gauss s * ?I s * ?I x \lborel \lborel)" by (auto simp: nn_integral_cmult[symmetric] nn_integral_multc[symmetric] * ennreal_mult[symmetric] intro!: nn_integral_cong split: split_indicator) also have "\ = (\\<^sup>+x. \\<^sup>+s. ?ff x s * ?I s * ?I x \lborel \lborel)" proof (rule nn_integral_cong, cases) fix x :: real assume "x \ 0" then show "(\\<^sup>+s. ?gauss x * ?gauss s * ?I s * ?I x \lborel) = (\\<^sup>+s. ?ff x s * ?I s * ?I x \lborel)" by (subst nn_integral_real_affine[where t="0" and c="x"]) (auto simp: mult_exp_exp nn_integral_cmult[symmetric] field_simps zero_less_mult_iff ennreal_mult[symmetric] intro!: nn_integral_cong split: split_indicator) qed simp also have "... = \\<^sup>+s. \\<^sup>+x. ?ff x s * ?I s * ?I x \lborel \lborel" by (rule lborel_pair.Fubini'[symmetric]) auto also have "... = ?pI (\s. ?pI (\x. ?ff x s))" by (rule nn_integral_cong_AE) (auto intro!: nn_integral_cong_AE AE_I[where N="{0}"] split: split_indicator_asm) also have "\ = ?pI (\s. ennreal (1 / (2 * (1 + s\<^sup>2))))" proof (intro nn_integral_cong ennreal_mult_right_cong) fix s :: real show "?pI (\x. ?ff x s) = ennreal (1 / (2 * (1 + s\<^sup>2)))" proof (subst nn_integral_FTC_atLeast) have "((\a. - (exp (- (a\<^sup>2 * (1 + s\<^sup>2))) / (2 + 2 * s\<^sup>2))) \ (- (0 / (2 + 2 * s\<^sup>2)))) at_top" apply (intro tendsto_intros filterlim_compose[OF exp_at_bot] filterlim_compose[OF filterlim_uminus_at_bot_at_top]) apply (subst mult.commute) apply (auto intro!: filterlim_tendsto_pos_mult_at_top filterlim_at_top_mult_at_top[OF filterlim_ident filterlim_ident] simp: add_pos_nonneg power2_eq_square add_nonneg_eq_0_iff) done then show "((\a. - (exp (- a\<^sup>2 - s\<^sup>2 * a\<^sup>2) / (2 + 2 * s\<^sup>2))) \ 0) at_top" by (simp add: field_simps) qed (auto intro!: derivative_eq_intros simp: field_simps add_nonneg_eq_0_iff) qed also have "... = ennreal (pi / 4)" proof (subst nn_integral_FTC_atLeast) show "((\a. arctan a / 2) \ (pi / 2) / 2 ) at_top" by (intro tendsto_intros) (simp_all add: tendsto_arctan_at_top) qed (auto intro!: derivative_eq_intros simp: add_nonneg_eq_0_iff field_simps power2_eq_square) finally have "?pI ?gauss^2 = pi / 4" by (simp add: power2_eq_square) then have "?pI ?gauss = sqrt (pi / 4)" using power_eq_iff_eq_base[of 2 "enn2real (?pI ?gauss)" "sqrt (pi / 4)"] by (cases "?pI ?gauss") (auto simp: power2_eq_square ennreal_mult[symmetric] ennreal_top_mult) also have "?pI ?gauss = (\\<^sup>+x. indicator {0..} x *\<^sub>R exp (- x\<^sup>2) \lborel)" by (intro nn_integral_cong) (simp split: split_indicator) also have "sqrt (pi / 4) = sqrt pi / 2" by (simp add: real_sqrt_divide) finally show ?thesis by (rule has_bochner_integral_nn_integral[rotated 3]) auto qed lemma gaussian_moment_1: "has_bochner_integral lborel (\x::real. indicator {0..} x *\<^sub>R (exp (- x\<^sup>2) * x)) (1 / 2)" proof - have "(\\<^sup>+x. indicator {0..} x *\<^sub>R (exp (- x\<^sup>2) * x) \lborel) = (\\<^sup>+x. ennreal (x * exp (- x\<^sup>2)) * indicator {0..} x \lborel)" by (intro nn_integral_cong) (auto simp: ac_simps split: split_indicator) also have "\ = ennreal (0 - (- exp (- 0\<^sup>2) / 2))" proof (rule nn_integral_FTC_atLeast) have "((\x::real. - exp (- x\<^sup>2) / 2) \ - 0 / 2) at_top" by (intro tendsto_divide tendsto_minus filterlim_compose[OF exp_at_bot] filterlim_compose[OF filterlim_uminus_at_bot_at_top] filterlim_pow_at_top filterlim_ident) auto then show "((\a::real. - exp (- a\<^sup>2) / 2) \ 0) at_top" by simp qed (auto intro!: derivative_eq_intros) also have "\ = ennreal (1 / 2)" by simp finally show ?thesis by (rule has_bochner_integral_nn_integral[rotated 3]) (auto split: split_indicator) qed lemma fixes k :: nat shows gaussian_moment_even_pos: "has_bochner_integral lborel (\x::real. indicator {0..} x *\<^sub>R (exp (-x\<^sup>2)*x^(2 * k))) ((sqrt pi / 2) * (fact (2 * k) / (2 ^ (2 * k) * fact k)))" (is "?even") and gaussian_moment_odd_pos: "has_bochner_integral lborel (\x::real. indicator {0..} x *\<^sub>R (exp (-x\<^sup>2)*x^(2 * k + 1))) (fact k / 2)" (is "?odd") proof - let ?M = "\k x. exp (- x\<^sup>2) * x^k :: real" { fix k I assume Mk: "has_bochner_integral lborel (\x. indicator {0..} x *\<^sub>R ?M k x) I" have "2 \ (0::real)" by linarith let ?f = "\b. \x. indicator {0..} x *\<^sub>R ?M (k + 2) x * indicator {..b} x \lborel" have "((\b. (k + 1) / 2 * (\x. indicator {..b} x *\<^sub>R (indicator {0..} x *\<^sub>R ?M k x) \lborel) - ?M (k + 1) b / 2) \ (k + 1) / 2 * (\x. indicator {0..} x *\<^sub>R ?M k x \lborel) - 0 / 2) at_top" (is ?tendsto) proof (intro tendsto_intros \2 \ 0\ tendsto_integral_at_top sets_lborel Mk[THEN integrable.intros]) show "(?M (k + 1) \ 0) at_top" proof cases assume "even k" have "((\x. ((x\<^sup>2)^(k div 2 + 1) / exp (x\<^sup>2)) * (1 / x) :: real) \ 0 * 0) at_top" by (intro tendsto_intros tendsto_divide_0[OF tendsto_const] filterlim_compose[OF tendsto_power_div_exp_0] filterlim_at_top_imp_at_infinity filterlim_ident filterlim_pow_at_top filterlim_ident) auto also have "(\x. ((x\<^sup>2)^(k div 2 + 1) / exp (x\<^sup>2)) * (1 / x) :: real) = ?M (k + 1)" using \even k\ by (auto simp: fun_eq_iff exp_minus field_simps power2_eq_square power_mult elim: evenE) finally show ?thesis by simp next assume "odd k" have "((\x. ((x\<^sup>2)^((k - 1) div 2 + 1) / exp (x\<^sup>2)) :: real) \ 0) at_top" by (intro filterlim_compose[OF tendsto_power_div_exp_0] filterlim_at_top_imp_at_infinity filterlim_ident filterlim_pow_at_top) auto also have "(\x. ((x\<^sup>2)^((k - 1) div 2 + 1) / exp (x\<^sup>2)) :: real) = ?M (k + 1)" using \odd k\ by (auto simp: fun_eq_iff exp_minus field_simps power2_eq_square power_mult elim: oddE) finally show ?thesis by simp qed qed also have "?tendsto \ ((?f \ (k + 1) / 2 * (\x. indicator {0..} x *\<^sub>R ?M k x \lborel) - 0 / 2) at_top)" proof (intro filterlim_cong refl eventually_at_top_linorder[THEN iffD2] exI[of _ 0] allI impI) fix b :: real assume b: "0 \ b" have "Suc k * (\x. indicator {0..b} x *\<^sub>R ?M k x \lborel) = (\x. indicator {0..b} x *\<^sub>R (exp (- x\<^sup>2) * ((Suc k) * x ^ k)) \lborel)" unfolding integral_mult_right_zero[symmetric] by (intro Bochner_Integration.integral_cong) auto also have "\ = exp (- b\<^sup>2) * b ^ (Suc k) - exp (- 0\<^sup>2) * 0 ^ (Suc k) - (\x. indicator {0..b} x *\<^sub>R (- 2 * x * exp (- x\<^sup>2) * x ^ (Suc k)) \lborel)" by (rule integral_by_parts') (auto intro!: derivative_eq_intros b simp: diff_Suc of_nat_Suc field_simps split: nat.split) also have "... = exp (- b\<^sup>2) * b ^ (Suc k) - (\x. indicator {0..b} x *\<^sub>R (- 2 * (exp (- x\<^sup>2) * x ^ (k + 2))) \lborel)" by (auto simp: intro!: Bochner_Integration.integral_cong) also have "... = exp (- b\<^sup>2) * b ^ (Suc k) + 2 * (\x. indicator {0..b} x *\<^sub>R ?M (k + 2) x \lborel)" unfolding Bochner_Integration.integral_mult_right_zero [symmetric] by (simp del: real_scaleR_def) finally have "Suc k * (\x. indicator {0..b} x *\<^sub>R ?M k x \lborel) = exp (- b\<^sup>2) * b ^ (Suc k) + 2 * (\x. indicator {0..b} x *\<^sub>R ?M (k + 2) x \lborel)" . then show "(k + 1) / 2 * (\x. indicator {..b} x *\<^sub>R (indicator {0..} x *\<^sub>R ?M k x)\lborel) - exp (- b\<^sup>2) * b ^ (k + 1) / 2 = ?f b" by (simp add: field_simps atLeastAtMost_def indicator_inter_arith) qed finally have int_M_at_top: "((?f \ (k + 1) / 2 * (\x. indicator {0..} x *\<^sub>R ?M k x \lborel)) at_top)" by simp have "has_bochner_integral lborel (\x. indicator {0..} x *\<^sub>R ?M (k + 2) x) ((k + 1) / 2 * I)" proof (rule has_bochner_integral_monotone_convergence_at_top) fix y :: real have *: "(\x. indicator {0..} x *\<^sub>R ?M (k + 2) x * indicator {..y} x::real) = (\x. indicator {0..y} x *\<^sub>R ?M (k + 2) x)" by rule (simp split: split_indicator) show "integrable lborel (\x. indicator {0..} x *\<^sub>R (?M (k + 2) x) * indicator {..y} x::real)" unfolding * by (rule borel_integrable_compact) (auto intro!: continuous_intros) show "((?f \ (k + 1) / 2 * I) at_top)" using int_M_at_top has_bochner_integral_integral_eq[OF Mk] by simp qed (auto split: split_indicator) } note step = this show ?even proof (induct k) case (Suc k) note step[OF this] also have "(real (2 * k + 1) / 2 * (sqrt pi / 2 * ((fact (2 * k)) / ((2::real)^(2*k) * fact k)))) = sqrt pi / 2 * ((fact (2 * Suc k)) / ((2::real)^(2 * Suc k) * fact (Suc k)))" apply (simp add: field_simps del: fact_Suc) apply (simp add: of_nat_mult field_simps) done finally show ?case by simp qed (insert gaussian_moment_0, simp) show ?odd proof (induct k) case (Suc k) note step[OF this] also have "(real (2 * k + 1 + 1) / (2::real) * ((fact k) / 2)) = (fact (Suc k)) / 2" by (simp add: field_simps of_nat_Suc field_split_simps del: fact_Suc) (simp add: field_simps) finally show ?case by simp qed (insert gaussian_moment_1, simp) qed context fixes k :: nat and \ \ :: real assumes [arith]: "0 < \" begin lemma normal_moment_even: "has_bochner_integral lborel (\x. normal_density \ \ x * (x - \) ^ (2 * k)) (fact (2 * k) / ((2 / \\<^sup>2)^k * fact k))" proof - have eq: "\x::real. x\<^sup>2^k = (x^k)\<^sup>2" by (simp add: power_mult[symmetric] ac_simps) have "has_bochner_integral lborel (\x. exp (-x\<^sup>2)*x^(2 * k)) (sqrt pi * (fact (2 * k) / (2 ^ (2 * k) * fact k)))" using has_bochner_integral_even_function[OF gaussian_moment_even_pos[where k=k]] by simp then have "has_bochner_integral lborel (\x. (exp (-x\<^sup>2)*x^(2 * k)) * ((2*\\<^sup>2)^k / sqrt (2 * pi * \\<^sup>2))) ((sqrt pi * (fact (2 * k) / (2 ^ (2 * k) * fact k))) * ((2*\\<^sup>2)^k / sqrt (2 * pi * \\<^sup>2)))" by (rule has_bochner_integral_mult_left) also have "(\x. (exp (-x\<^sup>2)*x^(2 * k)) * ((2*\\<^sup>2)^k / sqrt (2 * pi * \\<^sup>2))) = (\x. exp (- ((sqrt 2 * \) * x)\<^sup>2 / (2*\\<^sup>2)) * ((sqrt 2 * \) * x) ^ (2 * k) / sqrt (2 * pi * \\<^sup>2))" by (auto simp: fun_eq_iff field_simps real_sqrt_power[symmetric] real_sqrt_mult real_sqrt_divide power_mult eq) also have "((sqrt pi * (fact (2 * k) / (2 ^ (2 * k) * fact k))) * ((2*\\<^sup>2)^k / sqrt (2 * pi * \\<^sup>2))) = (inverse (sqrt 2 * \) * ((fact (2 * k))) / ((2/\\<^sup>2) ^ k * (fact k)))" by (auto simp: fun_eq_iff power_mult field_simps real_sqrt_power[symmetric] real_sqrt_mult power2_eq_square) finally show ?thesis unfolding normal_density_def by (subst lborel_has_bochner_integral_real_affine_iff[where c="sqrt 2 * \" and t=\]) simp_all qed lemma normal_moment_abs_odd: "has_bochner_integral lborel (\x. normal_density \ \ x * \x - \\^(2 * k + 1)) (2^k * \^(2 * k + 1) * fact k * sqrt (2 / pi))" proof - have "has_bochner_integral lborel (\x::real. indicator {0..} x *\<^sub>R (exp (-x\<^sup>2)*\x\^(2 * k + 1))) (fact k / 2)" by (rule has_bochner_integral_cong[THEN iffD1, OF _ _ _ gaussian_moment_odd_pos[of k]]) auto from has_bochner_integral_even_function[OF this] have "has_bochner_integral lborel (\x::real. exp (-x\<^sup>2)*\x\^(2 * k + 1)) (fact k)" by simp then have "has_bochner_integral lborel (\x. (exp (-x\<^sup>2)*\x\^(2 * k + 1)) * (2^k * \^(2 * k + 1) / sqrt (pi * \\<^sup>2))) (fact k * (2^k * \^(2 * k + 1) / sqrt (pi * \\<^sup>2)))" by (rule has_bochner_integral_mult_left) also have "(\x. (exp (-x\<^sup>2)*\x\^(2 * k + 1)) * (2^k * \^(2 * k + 1) / sqrt (pi * \\<^sup>2))) = (\x. exp (- (((sqrt 2 * \) * x)\<^sup>2 / (2 * \\<^sup>2))) * \sqrt 2 * \ * x\ ^ (2 * k + 1) / sqrt (2 * pi * \\<^sup>2))" by (simp add: field_simps abs_mult real_sqrt_power[symmetric] power_mult real_sqrt_mult) also have "(fact k * (2^k * \^(2 * k + 1) / sqrt (pi * \\<^sup>2))) = (inverse (sqrt 2) * inverse \ * (2 ^ k * (\ * \ ^ (2 * k)) * (fact k) * sqrt (2 / pi)))" by (auto simp: fun_eq_iff power_mult field_simps real_sqrt_power[symmetric] real_sqrt_divide real_sqrt_mult) finally show ?thesis unfolding normal_density_def by (subst lborel_has_bochner_integral_real_affine_iff[where c="sqrt 2 * \" and t=\]) simp_all qed lemma normal_moment_odd: "has_bochner_integral lborel (\x. normal_density \ \ x * (x - \)^(2 * k + 1)) 0" proof - have "has_bochner_integral lborel (\x. exp (- x\<^sup>2) * x^(2 * k + 1)::real) 0" using gaussian_moment_odd_pos by (rule has_bochner_integral_odd_function) simp then have "has_bochner_integral lborel (\x. (exp (-x\<^sup>2)*x^(2 * k + 1)) * (2^k*\^(2*k)/sqrt pi)) (0 * (2^k*\^(2*k)/sqrt pi))" by (rule has_bochner_integral_mult_left) also have "(\x. (exp (-x\<^sup>2)*x^(2 * k + 1)) * (2^k*\^(2*k)/sqrt pi)) = (\x. exp (- ((sqrt 2 * \ * x)\<^sup>2 / (2 * \\<^sup>2))) * (sqrt 2 * \ * x * (sqrt 2 * \ * x) ^ (2 * k)) / sqrt (2 * pi * \\<^sup>2))" unfolding real_sqrt_mult by (simp add: field_simps abs_mult real_sqrt_power[symmetric] power_mult fun_eq_iff) finally show ?thesis unfolding normal_density_def by (subst lborel_has_bochner_integral_real_affine_iff[where c="sqrt 2 * \" and t=\]) simp_all qed lemma integral_normal_moment_even: "integral\<^sup>L lborel (\x. normal_density \ \ x * (x - \)^(2 * k)) = fact (2 * k) / ((2 / \\<^sup>2)^k * fact k)" using normal_moment_even by (rule has_bochner_integral_integral_eq) lemma integral_normal_moment_abs_odd: "integral\<^sup>L lborel (\x. normal_density \ \ x * \x - \\^(2 * k + 1)) = 2 ^ k * \ ^ (2 * k + 1) * fact k * sqrt (2 / pi)" using normal_moment_abs_odd by (rule has_bochner_integral_integral_eq) lemma integral_normal_moment_odd: "integral\<^sup>L lborel (\x. normal_density \ \ x * (x - \)^(2 * k + 1)) = 0" using normal_moment_odd by (rule has_bochner_integral_integral_eq) end context fixes \ :: real assumes \_pos[arith]: "0 < \" begin lemma normal_moment_nz_1: "has_bochner_integral lborel (\x. normal_density \ \ x * x) \" proof - note normal_moment_even[OF \_pos, of \ 0] note normal_moment_odd[OF \_pos, of \ 0] has_bochner_integral_mult_left[of \, OF this] note has_bochner_integral_add[OF this] then show ?thesis by (simp add: power2_eq_square field_simps) qed lemma integral_normal_moment_nz_1: "integral\<^sup>L lborel (\x. normal_density \ \ x * x) = \" using normal_moment_nz_1 by (rule has_bochner_integral_integral_eq) lemma integrable_normal_moment_nz_1: "integrable lborel (\x. normal_density \ \ x * x)" using normal_moment_nz_1 by rule lemma integrable_normal_moment: "integrable lborel (\x. normal_density \ \ x * (x - \)^k)" proof cases assume "even k" then show ?thesis using integrable.intros[OF normal_moment_even] by (auto elim: evenE) next assume "odd k" then show ?thesis using integrable.intros[OF normal_moment_odd] by (auto elim: oddE) qed lemma integrable_normal_moment_abs: "integrable lborel (\x. normal_density \ \ x * \x - \\^k)" proof cases assume "even k" then show ?thesis using integrable.intros[OF normal_moment_even] by (auto simp add: power_even_abs elim: evenE) next assume "odd k" then show ?thesis using integrable.intros[OF normal_moment_abs_odd] by (auto elim: oddE) qed lemma integrable_normal_density[simp, intro]: "integrable lborel (normal_density \ \)" using integrable_normal_moment[of \ 0] by simp lemma integral_normal_density[simp]: "(\x. normal_density \ \ x \lborel) = 1" using integral_normal_moment_even[of \ \ 0] by simp lemma prob_space_normal_density: "prob_space (density lborel (normal_density \ \))" proof qed (simp add: emeasure_density nn_integral_eq_integral normal_density_nonneg) end context fixes k :: nat begin lemma std_normal_moment_even: "has_bochner_integral lborel (\x. std_normal_density x * x ^ (2 * k)) (fact (2 * k) / (2^k * fact k))" using normal_moment_even[of 1 0 k] by simp lemma std_normal_moment_abs_odd: "has_bochner_integral lborel (\x. std_normal_density x * \x\^(2 * k + 1)) (sqrt (2/pi) * 2^k * fact k)" using normal_moment_abs_odd[of 1 0 k] by (simp add: ac_simps) lemma std_normal_moment_odd: "has_bochner_integral lborel (\x. std_normal_density x * x^(2 * k + 1)) 0" using normal_moment_odd[of 1 0 k] by simp lemma integral_std_normal_moment_even: "integral\<^sup>L lborel (\x. std_normal_density x * x^(2*k)) = fact (2 * k) / (2^k * fact k)" using std_normal_moment_even by (rule has_bochner_integral_integral_eq) lemma integral_std_normal_moment_abs_odd: "integral\<^sup>L lborel (\x. std_normal_density x * \x\^(2 * k + 1)) = sqrt (2 / pi) * 2^k * fact k" using std_normal_moment_abs_odd by (rule has_bochner_integral_integral_eq) lemma integral_std_normal_moment_odd: "integral\<^sup>L lborel (\x. std_normal_density x * x^(2 * k + 1)) = 0" using std_normal_moment_odd by (rule has_bochner_integral_integral_eq) lemma integrable_std_normal_moment_abs: "integrable lborel (\x. std_normal_density x * \x\^k)" using integrable_normal_moment_abs[of 1 0 k] by simp lemma integrable_std_normal_moment: "integrable lborel (\x. std_normal_density x * x^k)" using integrable_normal_moment[of 1 0 k] by simp end lemma (in prob_space) normal_density_affine: assumes X: "distributed M lborel X (normal_density \ \)" assumes [simp, arith]: "0 < \" "\ \ 0" shows "distributed M lborel (\x. \ + \ * X x) (normal_density (\ + \ * \) (\\\ * \))" proof - have eq: "\x. \\\ * normal_density (\ + \ * \) (\\\ * \) (x * \ + \) = normal_density \ \ x" by (simp add: normal_density_def real_sqrt_mult field_simps) (simp add: power2_eq_square field_simps) show ?thesis by (rule distributed_affineI[OF _ \\ \ 0\, where t=\]) (simp_all add: eq X ennreal_mult'[symmetric]) qed lemma (in prob_space) normal_standard_normal_convert: assumes pos_var[simp, arith]: "0 < \" shows "distributed M lborel X (normal_density \ \) = distributed M lborel (\x. (X x - \) / \) std_normal_density" proof auto assume "distributed M lborel X (\x. ennreal (normal_density \ \ x))" then have "distributed M lborel (\x. -\ / \ + (1/\) * X x) (\x. ennreal (normal_density (-\ / \ + (1/\)* \) (\1/\\ * \) x))" by(rule normal_density_affine) auto then show "distributed M lborel (\x. (X x - \) / \) (\x. ennreal (std_normal_density x))" by (simp add: diff_divide_distrib[symmetric] field_simps) next assume *: "distributed M lborel (\x. (X x - \) / \) (\x. ennreal (std_normal_density x))" have "distributed M lborel (\x. \ + \ * ((X x - \) / \)) (\x. ennreal (normal_density \ \\\ x))" using normal_density_affine[OF *, of \ \] by simp then show "distributed M lborel X (\x. ennreal (normal_density \ \ x))" by simp qed lemma conv_normal_density_zero_mean: assumes [simp, arith]: "0 < \" "0 < \" shows "(\x. \\<^sup>+y. ennreal (normal_density 0 \ (x - y) * normal_density 0 \ y) \lborel) = normal_density 0 (sqrt (\\<^sup>2 + \\<^sup>2))" (is "?LHS = ?RHS") proof - define \' \' where "\' = \\<^sup>2" and "\' = \\<^sup>2" then have [simp, arith]: "0 < \'" "0 < \'" by simp_all let ?\ = "sqrt ((\' * \') / (\' + \'))" have sqrt: "(sqrt (2 * pi * (\' + \')) * sqrt (2 * pi * (\' * \') / (\' + \'))) = (sqrt (2 * pi * \') * sqrt (2 * pi * \'))" by (subst power_eq_iff_eq_base[symmetric, where n=2]) (simp_all add: real_sqrt_mult[symmetric] power2_eq_square) have "?LHS = (\x. \\<^sup>+y. ennreal((normal_density 0 (sqrt (\' + \')) x) * normal_density (\' * x / (\' + \')) ?\ y) \lborel)" apply (intro ext nn_integral_cong) apply (simp add: normal_density_def \'_def[symmetric] \'_def[symmetric] sqrt mult_exp_exp) apply (simp add: divide_simps power2_eq_square) apply (simp add: algebra_simps) done also have "... = (\x. (normal_density 0 (sqrt (\\<^sup>2 + \\<^sup>2)) x) * \\<^sup>+y. ennreal( normal_density (\\<^sup>2* x / (\\<^sup>2 + \\<^sup>2)) ?\ y) \lborel)" by (subst nn_integral_cmult[symmetric]) (auto simp: \'_def \'_def normal_density_def ennreal_mult'[symmetric]) also have "... = (\x. (normal_density 0 (sqrt (\\<^sup>2 + \\<^sup>2)) x))" by (subst nn_integral_eq_integral) (auto simp: normal_density_nonneg) finally show ?thesis by fast qed lemma conv_std_normal_density: "(\x. \\<^sup>+y. ennreal (std_normal_density (x - y) * std_normal_density y) \lborel) = (normal_density 0 (sqrt 2))" by (subst conv_normal_density_zero_mean) simp_all lemma (in prob_space) add_indep_normal: assumes indep: "indep_var borel X borel Y" assumes pos_var[arith]: "0 < \" "0 < \" assumes normalX[simp]: "distributed M lborel X (normal_density \ \)" assumes normalY[simp]: "distributed M lborel Y (normal_density \ \)" shows "distributed M lborel (\x. X x + Y x) (normal_density (\ + \) (sqrt (\\<^sup>2 + \\<^sup>2)))" proof - have ind[simp]: "indep_var borel (\x. - \ + X x) borel (\x. - \ + Y x)" proof - have "indep_var borel ( (\x. -\ + x) o X) borel ((\x. - \ + x) o Y)" by (auto intro!: indep_var_compose assms) then show ?thesis by (simp add: o_def) qed have "distributed M lborel (\x. -\ + 1 * X x) (normal_density (- \ + 1 * \) (\1\ * \))" by(rule normal_density_affine[OF normalX pos_var(1), of 1 "-\"]) simp then have 1[simp]: "distributed M lborel (\x. - \ + X x) (normal_density 0 \)" by simp have "distributed M lborel (\x. -\ + 1 * Y x) (normal_density (- \ + 1 * \) (\1\ * \))" by(rule normal_density_affine[OF normalY pos_var(2), of 1 "-\"]) simp then have 2[simp]: "distributed M lborel (\x. - \ + Y x) (normal_density 0 \)" by simp have *: "distributed M lborel (\x. (- \ + X x) + (- \ + Y x)) (\x. ennreal (normal_density 0 (sqrt (\\<^sup>2 + \\<^sup>2)) x))" using distributed_convolution[OF ind 1 2] conv_normal_density_zero_mean[OF pos_var] by (simp add: ennreal_mult'[symmetric] normal_density_nonneg) have "distributed M lborel (\x. \ + \ + 1 * (- \ + X x + (- \ + Y x))) (\x. ennreal (normal_density (\ + \ + 1 * 0) (\1\ * sqrt (\\<^sup>2 + \\<^sup>2)) x))" by (rule normal_density_affine[OF *, of 1 "\ + \"]) (auto simp: add_pos_pos) then show ?thesis by auto qed lemma (in prob_space) diff_indep_normal: assumes indep[simp]: "indep_var borel X borel Y" assumes [simp, arith]: "0 < \" "0 < \" assumes normalX[simp]: "distributed M lborel X (normal_density \ \)" assumes normalY[simp]: "distributed M lborel Y (normal_density \ \)" shows "distributed M lborel (\x. X x - Y x) (normal_density (\ - \) (sqrt (\\<^sup>2 + \\<^sup>2)))" proof - have "distributed M lborel (\x. 0 + - 1 * Y x) (\x. ennreal (normal_density (0 + - 1 * \) (\- 1\ * \) x))" by(rule normal_density_affine, auto) then have [simp]:"distributed M lborel (\x. - Y x) (\x. ennreal (normal_density (- \) \ x))" by simp have "distributed M lborel (\x. X x + (- Y x)) (normal_density (\ + - \) (sqrt (\\<^sup>2 + \\<^sup>2)))" apply (rule add_indep_normal) apply (rule indep_var_compose[unfolded comp_def, of borel _ borel _ "\x. x" _ "\x. - x"]) apply simp_all done then show ?thesis by simp qed lemma (in prob_space) sum_indep_normal: assumes "finite I" "I \ {}" "indep_vars (\i. borel) X I" assumes "\i. i \ I \ 0 < \ i" assumes normal: "\i. i \ I \ distributed M lborel (X i) (normal_density (\ i) (\ i))" shows "distributed M lborel (\x. \i\I. X i x) (normal_density (\i\I. \ i) (sqrt (\i\I. (\ i)\<^sup>2)))" using assms proof (induct I rule: finite_ne_induct) case (singleton i) then show ?case by (simp add : abs_of_pos) next case (insert i I) then have 1: "distributed M lborel (\x. (X i x) + (\i\I. X i x)) (normal_density (\ i + sum \ I) (sqrt ((\ i)\<^sup>2 + (sqrt (\i\I. (\ i)\<^sup>2))\<^sup>2)))" apply (intro add_indep_normal indep_vars_sum insert real_sqrt_gt_zero) apply (auto intro: indep_vars_subset intro!: sum_pos) apply fastforce done have 2: "(\x. (X i x)+ (\i\I. X i x)) = (\x. (\j\insert i I. X j x))" "\ i + sum \ I = sum \ (insert i I)" using insert by auto have 3: "(sqrt ((\ i)\<^sup>2 + (sqrt (\i\I. (\ i)\<^sup>2))\<^sup>2)) = (sqrt (\i\insert i I. (\ i)\<^sup>2))" using insert by (simp add: sum_nonneg) show ?case using 1 2 3 by simp qed lemma (in prob_space) standard_normal_distributed_expectation: assumes D: "distributed M lborel X std_normal_density" shows "expectation X = 0" using integral_std_normal_moment_odd[of 0] distributed_integral[OF D, of "\x. x", symmetric] - by (auto simp: ) + by auto lemma (in prob_space) normal_distributed_expectation: assumes \[arith]: "0 < \" assumes D: "distributed M lborel X (normal_density \ \)" shows "expectation X = \" using integral_normal_moment_nz_1[OF \, of \] distributed_integral[OF D, of "\x. x", symmetric] by (auto simp: field_simps) lemma (in prob_space) normal_distributed_variance: fixes a b :: real assumes [simp, arith]: "0 < \" assumes D: "distributed M lborel X (normal_density \ \)" shows "variance X = \\<^sup>2" using integral_normal_moment_even[of \ \ 1] by (subst distributed_integral[OF D, symmetric]) (simp_all add: eval_nat_numeral normal_distributed_expectation[OF assms]) lemma (in prob_space) standard_normal_distributed_variance: "distributed M lborel X std_normal_density \ variance X = 1" using normal_distributed_variance[of 1 X 0] by simp lemma (in information_space) entropy_normal_density: assumes [arith]: "0 < \" assumes D: "distributed M lborel X (normal_density \ \)" shows "entropy b lborel X = log b (2 * pi * exp 1 * \\<^sup>2) / 2" proof - have "entropy b lborel X = - (\ x. normal_density \ \ x * log b (normal_density \ \ x) \lborel)" using D by (rule entropy_distr) simp also have "\ = - (\ x. normal_density \ \ x * (- ln (2 * pi * \\<^sup>2) - (x - \)\<^sup>2 / \\<^sup>2) / (2 * ln b) \lborel)" by (intro arg_cong[where f="uminus"] Bochner_Integration.integral_cong) (auto simp: normal_density_def field_simps ln_mult log_def ln_div ln_sqrt) also have "\ = - (\x. - (normal_density \ \ x * (ln (2 * pi * \\<^sup>2)) + (normal_density \ \ x * (x - \)\<^sup>2) / \\<^sup>2) / (2 * ln b) \lborel)" by (intro arg_cong[where f="uminus"] Bochner_Integration.integral_cong) (auto simp: field_split_simps field_simps) also have "\ = (\x. normal_density \ \ x * (ln (2 * pi * \\<^sup>2)) + (normal_density \ \ x * (x - \)\<^sup>2) / \\<^sup>2 \lborel) / (2 * ln b)" by (simp del: minus_add_distrib) also have "\ = (ln (2 * pi * \\<^sup>2) + 1) / (2 * ln b)" using integral_normal_moment_even[of \ \ 1] by (simp add: integrable_normal_moment fact_numeral) also have "\ = log b (2 * pi * exp 1 * \\<^sup>2) / 2" by (simp add: log_def field_simps ln_mult) finally show ?thesis . qed end diff --git a/src/HOL/Probability/Product_PMF.thy b/src/HOL/Probability/Product_PMF.thy --- a/src/HOL/Probability/Product_PMF.thy +++ b/src/HOL/Probability/Product_PMF.thy @@ -1,910 +1,910 @@ (* File: Product_PMF.thy Authors: Manuel Eberl, Max W. Haslbeck *) section \Indexed products of PMFs\ theory Product_PMF imports Probability_Mass_Function Independent_Family begin subsection \Preliminaries\ lemma pmf_expectation_eq_infsetsum: "measure_pmf.expectation p f = infsetsum (\x. pmf p x * f x) UNIV" unfolding infsetsum_def measure_pmf_eq_density by (subst integral_density) simp_all lemma measure_pmf_prob_product: assumes "countable A" "countable B" shows "measure_pmf.prob (pair_pmf M N) (A \ B) = measure_pmf.prob M A * measure_pmf.prob N B" proof - have "measure_pmf.prob (pair_pmf M N) (A \ B) = (\\<^sub>a(a, b)\A \ B. pmf M a * pmf N b)" by (auto intro!: infsetsum_cong simp add: measure_pmf_conv_infsetsum pmf_pair) also have "\ = measure_pmf.prob M A * measure_pmf.prob N B" using assms by (subst infsetsum_product) (auto simp add: measure_pmf_conv_infsetsum) finally show ?thesis by simp qed subsection \Definition\ text \ In analogy to @{const PiM}, we define an indexed product of PMFs. In the literature, this is typically called taking a vector of independent random variables. Note that the components do not have to be identically distributed. The operation takes an explicit index set \<^term>\A :: 'a set\ and a function \<^term>\f :: 'a \ 'b pmf\ that maps each element from \<^term>\A\ to a PMF and defines the product measure $\bigotimes_{i\in A} f(i)$ , which is represented as a \<^typ>\('a \ 'b) pmf\. Note that unlike @{const PiM}, this only works for \<^emph>\finite\ index sets. It could be extended to countable sets and beyond, but the construction becomes somewhat more involved. \ definition Pi_pmf :: "'a set \ 'b \ ('a \ 'b pmf) \ ('a \ 'b) pmf" where "Pi_pmf A dflt p = embed_pmf (\f. if (\x. x \ A \ f x = dflt) then \x\A. pmf (p x) (f x) else 0)" text \ A technical subtlety that needs to be addressed is this: Intuitively, the functions in the support of a product distribution have domain \A\. However, since HOL is a total logic, these functions must still return \<^emph>\some\ value for inputs outside \A\. The product measure @{const PiM} simply lets these functions return @{const undefined} in these cases. We chose a different solution here, which is to supply a default value \<^term>\dflt :: 'b\ that is returned in these cases. As one possible application, one could model the result of \n\ different independent coin tosses as @{term "Pi_pmf {0.._. bernoulli_pmf (1 / 2))"}. This returns a function of type \<^typ>\nat \ bool\ that maps every natural number below \n\ to the result of the corresponding coin toss, and every other natural number to \<^term>\False\. \ lemma pmf_Pi: assumes A: "finite A" shows "pmf (Pi_pmf A dflt p) f = (if (\x. x \ A \ f x = dflt) then \x\A. pmf (p x) (f x) else 0)" unfolding Pi_pmf_def proof (rule pmf_embed_pmf, goal_cases) case 2 define S where "S = {f. \x. x \ A \ f x = dflt}" define B where "B = (\x. set_pmf (p x))" have neutral_left: "(\xa\A. pmf (p xa) (f xa)) = 0" if "f \ PiE A B - (\f. restrict f A) ` S" for f proof - have "restrict (\x. if x \ A then f x else dflt) A \ (\f. restrict f A) ` S" by (intro imageI) (auto simp: S_def) also have "restrict (\x. if x \ A then f x else dflt) A = f" using that by (auto simp: PiE_def Pi_def extensional_def fun_eq_iff) finally show ?thesis using that by blast qed have neutral_right: "(\xa\A. pmf (p xa) (f xa)) = 0" if "f \ (\f. restrict f A) ` S - PiE A B" for f proof - from that obtain f' where f': "f = restrict f' A" "f' \ S" by auto moreover from this and that have "restrict f' A \ PiE A B" by simp then obtain x where "x \ A" "pmf (p x) (f' x) = 0" by (auto simp: B_def set_pmf_eq) with f' and A show ?thesis by auto qed have "(\f. \x\A. pmf (p x) (f x)) abs_summable_on PiE A B" by (intro abs_summable_on_prod_PiE A) (auto simp: B_def) also have "?this \ (\f. \x\A. pmf (p x) (f x)) abs_summable_on (\f. restrict f A) ` S" by (intro abs_summable_on_cong_neutral neutral_left neutral_right) auto also have "\ \ (\f. \x\A. pmf (p x) (restrict f A x)) abs_summable_on S" by (rule abs_summable_on_reindex_iff [symmetric]) (force simp: inj_on_def fun_eq_iff S_def) also have "\ \ (\f. if \x. x \ A \ f x = dflt then \x\A. pmf (p x) (f x) else 0) abs_summable_on UNIV" by (intro abs_summable_on_cong_neutral) (auto simp: S_def) finally have summable: \ . have "1 = (\x\A. 1::real)" by simp also have "(\x\A. 1) = (\x\A. \\<^sub>ay\B x. pmf (p x) y)" unfolding B_def by (subst infsetsum_pmf_eq_1) auto also have "(\x\A. \\<^sub>ay\B x. pmf (p x) y) = (\\<^sub>af\Pi\<^sub>E A B. \x\A. pmf (p x) (f x))" by (intro infsetsum_prod_PiE [symmetric] A) (auto simp: B_def) also have "\ = (\\<^sub>af\(\f. restrict f A) ` S. \x\A. pmf (p x) (f x))" using A by (intro infsetsum_cong_neutral neutral_left neutral_right refl) also have "\ = (\\<^sub>af\S. \x\A. pmf (p x) (restrict f A x))" by (rule infsetsum_reindex) (force simp: inj_on_def fun_eq_iff S_def) also have "\ = (\\<^sub>af\S. \x\A. pmf (p x) (f x))" by (intro infsetsum_cong) (auto simp: S_def) also have "\ = (\\<^sub>af. if \x. x \ A \ f x = dflt then \x\A. pmf (p x) (f x) else 0)" by (intro infsetsum_cong_neutral) (auto simp: S_def) also have "ennreal \ = (\\<^sup>+f. ennreal (if \x. x \ A \ f x = dflt then \x\A. pmf (p x) (f x) else 0) \count_space UNIV)" by (intro nn_integral_conv_infsetsum [symmetric] summable) (auto simp: prod_nonneg) finally show ?case by simp qed (auto simp: prod_nonneg) lemma Pi_pmf_cong: assumes "A = A'" "dflt = dflt'" "\x. x \ A \ f x = f' x" shows "Pi_pmf A dflt f = Pi_pmf A' dflt' f'" proof - have "(\fa. if \x. x \ A \ fa x = dflt then \x\A. pmf (f x) (fa x) else 0) = (\f. if \x. x \ A' \ f x = dflt' then \x\A'. pmf (f' x) (f x) else 0)" using assms by (intro ext) (auto intro!: prod.cong) thus ?thesis by (simp only: Pi_pmf_def) qed lemma pmf_Pi': assumes "finite A" "\x. x \ A \ f x = dflt" shows "pmf (Pi_pmf A dflt p) f = (\x\A. pmf (p x) (f x))" using assms by (subst pmf_Pi) auto lemma pmf_Pi_outside: assumes "finite A" "\x. x \ A \ f x \ dflt" shows "pmf (Pi_pmf A dflt p) f = 0" using assms by (subst pmf_Pi) auto lemma pmf_Pi_empty [simp]: "Pi_pmf {} dflt p = return_pmf (\_. dflt)" by (intro pmf_eqI, subst pmf_Pi) (auto simp: indicator_def) lemma set_Pi_pmf_subset: "finite A \ set_pmf (Pi_pmf A dflt p) \ {f. \x. x \ A \ f x = dflt}" by (auto simp: set_pmf_eq pmf_Pi) subsection \Dependent product sets with a default\ text \ The following describes a dependent product of sets where the functions are required to return the default value \<^term>\dflt\ outside their domain, in analogy to @{const PiE}, which uses @{const undefined}. \ definition PiE_dflt where "PiE_dflt A dflt B = {f. \x. (x \ A \ f x \ B x) \ (x \ A \ f x = dflt)}" lemma restrict_PiE_dflt: "(\h. restrict h A) ` PiE_dflt A dflt B = PiE A B" proof (intro equalityI subsetI) fix h assume "h \ (\h. restrict h A) ` PiE_dflt A dflt B" thus "h \ PiE A B" by (auto simp: PiE_dflt_def) next fix h assume h: "h \ PiE A B" hence "restrict (\x. if x \ A then h x else dflt) A \ (\h. restrict h A) ` PiE_dflt A dflt B" by (intro imageI) (auto simp: PiE_def extensional_def PiE_dflt_def) also have "restrict (\x. if x \ A then h x else dflt) A = h" using h by (auto simp: fun_eq_iff) finally show "h \ (\h. restrict h A) ` PiE_dflt A dflt B" . qed lemma dflt_image_PiE: "(\h x. if x \ A then h x else dflt) ` PiE A B = PiE_dflt A dflt B" (is "?f ` ?X = ?Y") proof (intro equalityI subsetI) fix h assume "h \ ?f ` ?X" thus "h \ ?Y" by (auto simp: PiE_dflt_def PiE_def) next fix h assume h: "h \ ?Y" hence "?f (restrict h A) \ ?f ` ?X" by (intro imageI) (auto simp: PiE_def extensional_def PiE_dflt_def) also have "?f (restrict h A) = h" using h by (auto simp: fun_eq_iff PiE_dflt_def) finally show "h \ ?f ` ?X" . qed lemma finite_PiE_dflt [intro]: assumes "finite A" "\x. x \ A \ finite (B x)" shows "finite (PiE_dflt A d B)" proof - have "PiE_dflt A d B = (\f x. if x \ A then f x else d) ` PiE A B" by (rule dflt_image_PiE [symmetric]) also have "finite \" by (intro finite_imageI finite_PiE assms) finally show ?thesis . qed lemma card_PiE_dflt: assumes "finite A" "\x. x \ A \ finite (B x)" shows "card (PiE_dflt A d B) = (\x\A. card (B x))" proof - from assms have "(\x\A. card (B x)) = card (PiE A B)" by (intro card_PiE [symmetric]) auto also have "PiE A B = (\f. restrict f A) ` PiE_dflt A d B" by (rule restrict_PiE_dflt [symmetric]) also have "card \ = card (PiE_dflt A d B)" by (intro card_image) (force simp: inj_on_def restrict_def fun_eq_iff PiE_dflt_def) finally show ?thesis .. qed lemma PiE_dflt_empty_iff [simp]: "PiE_dflt A dflt B = {} \ (\x\A. B x = {})" by (simp add: dflt_image_PiE [symmetric] PiE_eq_empty_iff) lemma set_Pi_pmf_subset': assumes "finite A" shows "set_pmf (Pi_pmf A dflt p) \ PiE_dflt A dflt (set_pmf \ p)" using assms by (auto simp: set_pmf_eq pmf_Pi PiE_dflt_def) lemma set_Pi_pmf: assumes "finite A" shows "set_pmf (Pi_pmf A dflt p) = PiE_dflt A dflt (set_pmf \ p)" proof (rule equalityI) show "PiE_dflt A dflt (set_pmf \ p) \ set_pmf (Pi_pmf A dflt p)" proof safe fix f assume f: "f \ PiE_dflt A dflt (set_pmf \ p)" hence "pmf (Pi_pmf A dflt p) f = (\x\A. pmf (p x) (f x))" using assms by (auto simp: pmf_Pi PiE_dflt_def) also have "\ > 0" using f by (intro prod_pos) (auto simp: PiE_dflt_def set_pmf_eq) finally show "f \ set_pmf (Pi_pmf A dflt p)" by (auto simp: set_pmf_eq) qed qed (use set_Pi_pmf_subset'[OF assms, of dflt p] in auto) text \ The probability of an independent combination of events is precisely the product of the probabilities of each individual event. \ lemma measure_Pi_pmf_PiE_dflt: assumes [simp]: "finite A" shows "measure_pmf.prob (Pi_pmf A dflt p) (PiE_dflt A dflt B) = (\x\A. measure_pmf.prob (p x) (B x))" proof - define B' where "B' = (\x. B x \ set_pmf (p x))" have "measure_pmf.prob (Pi_pmf A dflt p) (PiE_dflt A dflt B) = (\\<^sub>ah\PiE_dflt A dflt B. pmf (Pi_pmf A dflt p) h)" by (rule measure_pmf_conv_infsetsum) also have "\ = (\\<^sub>ah\PiE_dflt A dflt B. \x\A. pmf (p x) (h x))" by (intro infsetsum_cong, subst pmf_Pi') (auto simp: PiE_dflt_def) also have "\ = (\\<^sub>ah\(\h. restrict h A) ` PiE_dflt A dflt B. \x\A. pmf (p x) (h x))" by (subst infsetsum_reindex) (force simp: inj_on_def PiE_dflt_def fun_eq_iff)+ also have "(\h. restrict h A) ` PiE_dflt A dflt B = PiE A B" by (rule restrict_PiE_dflt) also have "(\\<^sub>ah\PiE A B. \x\A. pmf (p x) (h x)) = (\\<^sub>ah\PiE A B'. \x\A. pmf (p x) (h x))" by (intro infsetsum_cong_neutral) (auto simp: B'_def set_pmf_eq) also have "(\\<^sub>ah\PiE A B'. \x\A. pmf (p x) (h x)) = (\x\A. infsetsum (pmf (p x)) (B' x))" by (intro infsetsum_prod_PiE) (auto simp: B'_def) also have "\ = (\x\A. infsetsum (pmf (p x)) (B x))" by (intro prod.cong infsetsum_cong_neutral) (auto simp: B'_def set_pmf_eq) also have "\ = (\x\A. measure_pmf.prob (p x) (B x))" by (subst measure_pmf_conv_infsetsum) (rule refl) finally show ?thesis . qed lemma measure_Pi_pmf_Pi: fixes t::nat assumes [simp]: "finite A" shows "measure_pmf.prob (Pi_pmf A dflt p) (Pi A B) = (\x\A. measure_pmf.prob (p x) (B x))" (is "?lhs = ?rhs") proof - have "?lhs = measure_pmf.prob (Pi_pmf A dflt p) (PiE_dflt A dflt B)" by (intro measure_prob_cong_0) (auto simp: PiE_dflt_def PiE_def intro!: pmf_Pi_outside)+ also have "\ = ?rhs" using assms by (simp add: measure_Pi_pmf_PiE_dflt) finally show ?thesis by simp qed subsection \Common PMF operations on products\ text \ @{const Pi_pmf} distributes over the `bind' operation in the Giry monad: \ lemma Pi_pmf_bind: assumes "finite A" shows "Pi_pmf A d (\x. bind_pmf (p x) (q x)) = do {f \ Pi_pmf A d' p; Pi_pmf A d (\x. q x (f x))}" (is "?lhs = ?rhs") proof (rule pmf_eqI, goal_cases) case (1 f) show ?case proof (cases "\x\-A. f x \ d") case False define B where "B = (\x. set_pmf (p x))" have [simp]: "countable (B x)" for x by (auto simp: B_def) { fix x :: 'a have "(\a. pmf (p x) a * 1) abs_summable_on B x" by (simp add: pmf_abs_summable) moreover have "norm (pmf (p x) a * 1) \ norm (pmf (p x) a * pmf (q x a) (f x))" for a unfolding norm_mult by (intro mult_left_mono) (auto simp: pmf_le_1) ultimately have "(\a. pmf (p x) a * pmf (q x a) (f x)) abs_summable_on B x" by (rule abs_summable_on_comparison_test) } note summable = this have "pmf ?rhs f = (\\<^sub>ag. pmf (Pi_pmf A d' p) g * (\x\A. pmf (q x (g x)) (f x)))" by (subst pmf_bind, subst pmf_Pi') (insert assms False, simp_all add: pmf_expectation_eq_infsetsum) also have "\ = (\\<^sub>ag\PiE_dflt A d' B. pmf (Pi_pmf A d' p) g * (\x\A. pmf (q x (g x)) (f x)))" unfolding B_def using assms by (intro infsetsum_cong_neutral) (auto simp: pmf_Pi PiE_dflt_def set_pmf_eq) also have "\ = (\\<^sub>ag\PiE_dflt A d' B. (\x\A. pmf (p x) (g x) * pmf (q x (g x)) (f x)))" using assms by (intro infsetsum_cong) (auto simp: pmf_Pi PiE_dflt_def prod.distrib) also have "\ = (\\<^sub>ag\(\g. restrict g A) ` PiE_dflt A d' B. (\x\A. pmf (p x) (g x) * pmf (q x (g x)) (f x)))" by (subst infsetsum_reindex) (force simp: PiE_dflt_def inj_on_def fun_eq_iff)+ also have "(\g. restrict g A) ` PiE_dflt A d' B = PiE A B" by (rule restrict_PiE_dflt) also have "(\\<^sub>ag\\. (\x\A. pmf (p x) (g x) * pmf (q x (g x)) (f x))) = (\x\A. \\<^sub>aa\B x. pmf (p x) a * pmf (q x a) (f x))" using assms summable by (subst infsetsum_prod_PiE) simp_all also have "\ = (\x\A. \\<^sub>aa. pmf (p x) a * pmf (q x a) (f x))" by (intro prod.cong infsetsum_cong_neutral) (auto simp: B_def set_pmf_eq) also have "\ = pmf ?lhs f" using False assms by (subst pmf_Pi') (simp_all add: pmf_bind pmf_expectation_eq_infsetsum) finally show ?thesis .. next case True have "pmf ?rhs f = measure_pmf.expectation (Pi_pmf A d' p) (\x. pmf (Pi_pmf A d (\xa. q xa (x xa))) f)" using assms by (simp add: pmf_bind) also have "\ = measure_pmf.expectation (Pi_pmf A d' p) (\x. 0)" using assms True by (intro Bochner_Integration.integral_cong pmf_Pi_outside) auto also have "\ = pmf ?lhs f" using assms True by (subst pmf_Pi_outside) auto finally show ?thesis .. qed qed lemma Pi_pmf_return_pmf [simp]: assumes "finite A" shows "Pi_pmf A dflt (\x. return_pmf (f x)) = return_pmf (\x. if x \ A then f x else dflt)" using assms by (intro pmf_eqI) (auto simp: pmf_Pi simp: indicator_def split: if_splits) text \ Analogously any componentwise mapping can be pulled outside the product: \ lemma Pi_pmf_map: assumes [simp]: "finite A" and "f dflt = dflt'" shows "Pi_pmf A dflt' (\x. map_pmf f (g x)) = map_pmf (\h. f \ h) (Pi_pmf A dflt g)" proof - have "Pi_pmf A dflt' (\x. map_pmf f (g x)) = Pi_pmf A dflt' (\x. g x \ (\x. return_pmf (f x)))" using assms by (simp add: map_pmf_def Pi_pmf_bind) also have "\ = Pi_pmf A dflt g \ (\h. return_pmf (\x. if x \ A then f (h x) else dflt'))" - by (subst Pi_pmf_bind[where d' = dflt]) (auto simp: ) + by (subst Pi_pmf_bind[where d' = dflt]) auto also have "\ = map_pmf (\h. f \ h) (Pi_pmf A dflt g)" unfolding map_pmf_def using set_Pi_pmf_subset'[of A dflt g] by (intro bind_pmf_cong refl arg_cong[of _ _ return_pmf]) (auto dest: simp: fun_eq_iff PiE_dflt_def assms(2)) finally show ?thesis . qed text \ We can exchange the default value in a product of PMFs like this: \ lemma Pi_pmf_default_swap: assumes "finite A" shows "map_pmf (\f x. if x \ A then f x else dflt') (Pi_pmf A dflt p) = Pi_pmf A dflt' p" (is "?lhs = ?rhs") proof (rule pmf_eqI, goal_cases) case (1 f) let ?B = "(\f x. if x \ A then f x else dflt') -` {f} \ PiE_dflt A dflt (\_. UNIV)" show ?case proof (cases "\x\-A. f x \ dflt'") case False let ?f' = "\x. if x \ A then f x else dflt" from False have "pmf ?lhs f = measure_pmf.prob (Pi_pmf A dflt p) ?B" using assms unfolding pmf_map by (intro measure_prob_cong_0) (auto simp: PiE_dflt_def pmf_Pi_outside) also from False have "?B = {?f'}" by (auto simp: fun_eq_iff PiE_dflt_def) also have "measure_pmf.prob (Pi_pmf A dflt p) {?f'} = pmf (Pi_pmf A dflt p) ?f'" by (simp add: measure_pmf_single) also have "\ = pmf ?rhs f" using False assms by (subst (1 2) pmf_Pi) auto finally show ?thesis . next case True have "pmf ?lhs f = measure_pmf.prob (Pi_pmf A dflt p) ?B" using assms unfolding pmf_map by (intro measure_prob_cong_0) (auto simp: PiE_dflt_def pmf_Pi_outside) also from True have "?B = {}" by auto also have "measure_pmf.prob (Pi_pmf A dflt p) \ = 0" by simp also have "0 = pmf ?rhs f" using True assms by (intro pmf_Pi_outside [symmetric]) auto finally show ?thesis . qed qed text \ The following rule allows reindexing the product: \ lemma Pi_pmf_bij_betw: assumes "finite A" "bij_betw h A B" "\x. x \ A \ h x \ B" shows "Pi_pmf A dflt (\_. f) = map_pmf (\g. g \ h) (Pi_pmf B dflt (\_. f))" (is "?lhs = ?rhs") proof - have B: "finite B" using assms bij_betw_finite by auto have "pmf ?lhs g = pmf ?rhs g" for g proof (cases "\a. a \ A \ g a = dflt") case True define h' where "h' = the_inv_into A h" have h': "h' (h x) = x" if "x \ A" for x unfolding h'_def using that assms by (auto simp add: bij_betw_def the_inv_into_f_f) have h: "h (h' x) = x" if "x \ B" for x unfolding h'_def using that assms f_the_inv_into_f_bij_betw by fastforce have "pmf ?rhs g = measure_pmf.prob (Pi_pmf B dflt (\_. f)) ((\g. g \ h) -` {g})" unfolding pmf_map by simp also have "\ = measure_pmf.prob (Pi_pmf B dflt (\_. f)) (((\g. g \ h) -` {g}) \ PiE_dflt B dflt (\_. UNIV))" using B by (intro measure_prob_cong_0) (auto simp: PiE_dflt_def pmf_Pi_outside) also have "\ = pmf (Pi_pmf B dflt (\_. f)) (\x. if x \ B then g (h' x) else dflt)" proof - have "(if h x \ B then g (h' (h x)) else dflt) = g x" for x using h' assms True by (cases "x \ A") (auto simp add: bij_betwE) then have "(\g. g \ h) -` {g} \ PiE_dflt B dflt (\_. UNIV) = {(\x. if x \ B then g (h' x) else dflt)}" using assms h' h True unfolding PiE_dflt_def by auto then show ?thesis by (simp add: measure_pmf_single) qed also have "\ = pmf (Pi_pmf A dflt (\_. f)) g" using B assms True h'_def by (auto simp add: pmf_Pi intro!: prod.reindex_bij_betw bij_betw_the_inv_into) finally show ?thesis by simp next case False have "pmf ?rhs g = infsetsum (pmf (Pi_pmf B dflt (\_. f))) ((\g. g \ h) -` {g})" using assms by (auto simp add: measure_pmf_conv_infsetsum pmf_map) also have "\ = infsetsum (\_. 0) ((\g x. g (h x)) -` {g})" using B False assms by (intro infsetsum_cong pmf_Pi_outside) fastforce+ also have "\ = 0" by simp finally show ?thesis using assms False by (auto simp add: pmf_Pi pmf_map) qed then show ?thesis by (rule pmf_eqI) qed text \ A product of uniform random choices is again a uniform distribution. \ lemma Pi_pmf_of_set: assumes "finite A" "\x. x \ A \ finite (B x)" "\x. x \ A \ B x \ {}" shows "Pi_pmf A d (\x. pmf_of_set (B x)) = pmf_of_set (PiE_dflt A d B)" (is "?lhs = ?rhs") proof (rule pmf_eqI, goal_cases) case (1 f) show ?case proof (cases "\x. x \ A \ f x \ d") case True hence "pmf ?lhs f = 0" using assms by (intro pmf_Pi_outside) (auto simp: PiE_dflt_def) also from True have "f \ PiE_dflt A d B" by (auto simp: PiE_dflt_def) hence "0 = pmf ?rhs f" using assms by (subst pmf_of_set) auto finally show ?thesis . next case False hence "pmf ?lhs f = (\x\A. pmf (pmf_of_set (B x)) (f x))" using assms by (subst pmf_Pi') auto also have "\ = (\x\A. indicator (B x) (f x) / real (card (B x)))" by (intro prod.cong refl, subst pmf_of_set) (use assms False in auto) also have "\ = (\x\A. indicator (B x) (f x)) / real (\x\A. card (B x))" by (subst prod_dividef) simp_all also have "(\x\A. indicator (B x) (f x) :: real) = indicator (PiE_dflt A d B) f" using assms False by (auto simp: indicator_def PiE_dflt_def) also have "(\x\A. card (B x)) = card (PiE_dflt A d B)" using assms by (intro card_PiE_dflt [symmetric]) auto also have "indicator (PiE_dflt A d B) f / \ = pmf ?rhs f" using assms by (intro pmf_of_set [symmetric]) auto finally show ?thesis . qed qed subsection \Merging and splitting PMF products\ text \ The following lemma shows that we can add a single PMF to a product: \ lemma Pi_pmf_insert: assumes "finite A" "x \ A" shows "Pi_pmf (insert x A) dflt p = map_pmf (\(y,f). f(x:=y)) (pair_pmf (p x) (Pi_pmf A dflt p))" proof (intro pmf_eqI) fix f let ?M = "pair_pmf (p x) (Pi_pmf A dflt p)" have "pmf (map_pmf (\(y, f). f(x := y)) ?M) f = measure_pmf.prob ?M ((\(y, f). f(x := y)) -` {f})" by (subst pmf_map) auto also have "((\(y, f). f(x := y)) -` {f}) = (\y'. {(f x, f(x := y'))})" by (auto simp: fun_upd_def fun_eq_iff) also have "measure_pmf.prob ?M \ = measure_pmf.prob ?M {(f x, f(x := dflt))}" using assms by (intro measure_prob_cong_0) (auto simp: pmf_pair pmf_Pi split: if_splits) also have "\ = pmf (p x) (f x) * pmf (Pi_pmf A dflt p) (f(x := dflt))" by (simp add: measure_pmf_single pmf_pair pmf_Pi) also have "\ = pmf (Pi_pmf (insert x A) dflt p) f" proof (cases "\y. y \ insert x A \ f y = dflt") case True with assms have "pmf (p x) (f x) * pmf (Pi_pmf A dflt p) (f(x := dflt)) = pmf (p x) (f x) * (\xa\A. pmf (p xa) ((f(x := dflt)) xa))" by (subst pmf_Pi') auto also have "(\xa\A. pmf (p xa) ((f(x := dflt)) xa)) = (\xa\A. pmf (p xa) (f xa))" using assms by (intro prod.cong) auto also have "pmf (p x) (f x) * \ = pmf (Pi_pmf (insert x A) dflt p) f" using assms True by (subst pmf_Pi') auto finally show ?thesis . qed (insert assms, auto simp: pmf_Pi) finally show "\ = pmf (map_pmf (\(y, f). f(x := y)) ?M) f" .. qed lemma Pi_pmf_insert': assumes "finite A" "x \ A" shows "Pi_pmf (insert x A) dflt p = do {y \ p x; f \ Pi_pmf A dflt p; return_pmf (f(x := y))}" using assms by (subst Pi_pmf_insert) (auto simp add: map_pmf_def pair_pmf_def case_prod_beta' bind_return_pmf bind_assoc_pmf) lemma Pi_pmf_singleton: "Pi_pmf {x} dflt p = map_pmf (\a b. if b = x then a else dflt) (p x)" proof - have "Pi_pmf {x} dflt p = map_pmf (fun_upd (\_. dflt) x) (p x)" by (subst Pi_pmf_insert) (simp_all add: pair_return_pmf2 pmf.map_comp o_def) also have "fun_upd (\_. dflt) x = (\z y. if y = x then z else dflt)" by (simp add: fun_upd_def fun_eq_iff) finally show ?thesis . qed text \ Projecting a product of PMFs onto a component yields the expected result: \ lemma Pi_pmf_component: assumes "finite A" shows "map_pmf (\f. f x) (Pi_pmf A dflt p) = (if x \ A then p x else return_pmf dflt)" proof (cases "x \ A") case True define A' where "A' = A - {x}" from assms and True have A': "A = insert x A'" by (auto simp: A'_def) from assms have "map_pmf (\f. f x) (Pi_pmf A dflt p) = p x" unfolding A' by (subst Pi_pmf_insert) (auto simp: A'_def pmf.map_comp o_def case_prod_unfold map_fst_pair_pmf) with True show ?thesis by simp next case False have "map_pmf (\f. f x) (Pi_pmf A dflt p) = map_pmf (\_. dflt) (Pi_pmf A dflt p)" using assms False set_Pi_pmf_subset[of A dflt p] by (intro pmf.map_cong refl) (auto simp: set_pmf_eq pmf_Pi_outside) with False show ?thesis by simp qed text \ We can take merge two PMF products on disjoint sets like this: \ lemma Pi_pmf_union: assumes "finite A" "finite B" "A \ B = {}" shows "Pi_pmf (A \ B) dflt p = map_pmf (\(f,g) x. if x \ A then f x else g x) (pair_pmf (Pi_pmf A dflt p) (Pi_pmf B dflt p))" (is "_ = map_pmf (?h A) (?q A)") using assms(1,3) proof (induction rule: finite_induct) case (insert x A) have "map_pmf (?h (insert x A)) (?q (insert x A)) = do {v \ p x; (f, g) \ pair_pmf (Pi_pmf A dflt p) (Pi_pmf B dflt p); return_pmf (\y. if y \ insert x A then (f(x := v)) y else g y)}" by (subst Pi_pmf_insert) (insert insert.hyps insert.prems, simp_all add: pair_pmf_def map_bind_pmf bind_map_pmf bind_assoc_pmf bind_return_pmf) also have "\ = do {v \ p x; (f, g) \ ?q A; return_pmf ((?h A (f,g))(x := v))}" by (intro bind_pmf_cong refl) (auto simp: fun_eq_iff) also have "\ = do {v \ p x; f \ map_pmf (?h A) (?q A); return_pmf (f(x := v))}" by (simp add: bind_map_pmf map_bind_pmf case_prod_unfold cong: if_cong) also have "\ = do {v \ p x; f \ Pi_pmf (A \ B) dflt p; return_pmf (f(x := v))}" using insert.hyps and insert.prems by (intro bind_pmf_cong insert.IH [symmetric] refl) auto also have "\ = Pi_pmf (insert x (A \ B)) dflt p" by (subst Pi_pmf_insert) (insert assms insert.hyps insert.prems, auto simp: pair_pmf_def map_bind_pmf) also have "insert x (A \ B) = insert x A \ B" by simp finally show ?case .. qed (simp_all add: case_prod_unfold map_snd_pair_pmf) text \ We can also project a product to a subset of the indices by mapping all the other indices to the default value: \ lemma Pi_pmf_subset: assumes "finite A" "A' \ A" shows "Pi_pmf A' dflt p = map_pmf (\f x. if x \ A' then f x else dflt) (Pi_pmf A dflt p)" proof - let ?P = "pair_pmf (Pi_pmf A' dflt p) (Pi_pmf (A - A') dflt p)" from assms have [simp]: "finite A'" by (blast dest: finite_subset) from assms have "A = A' \ (A - A')" by blast also have "Pi_pmf \ dflt p = map_pmf (\(f,g) x. if x \ A' then f x else g x) ?P" using assms by (intro Pi_pmf_union) auto also have "map_pmf (\f x. if x \ A' then f x else dflt) \ = map_pmf fst ?P" unfolding map_pmf_comp o_def case_prod_unfold using set_Pi_pmf_subset[of A' dflt p] by (intro map_pmf_cong refl) (auto simp: fun_eq_iff) also have "\ = Pi_pmf A' dflt p" by (simp add: map_fst_pair_pmf) finally show ?thesis .. qed lemma Pi_pmf_subset': fixes f :: "'a \ 'b pmf" assumes "finite A" "B \ A" "\x. x \ A - B \ f x = return_pmf dflt" shows "Pi_pmf A dflt f = Pi_pmf B dflt f" proof - have "Pi_pmf (B \ (A - B)) dflt f = map_pmf (\(f, g) x. if x \ B then f x else g x) (pair_pmf (Pi_pmf B dflt f) (Pi_pmf (A - B) dflt f))" using assms by (intro Pi_pmf_union) (auto dest: finite_subset) also have "Pi_pmf (A - B) dflt f = Pi_pmf (A - B) dflt (\_. return_pmf dflt)" using assms by (intro Pi_pmf_cong) auto also have "\ = return_pmf (\_. dflt)" using assms by simp also have "map_pmf (\(f, g) x. if x \ B then f x else g x) (pair_pmf (Pi_pmf B dflt f) (return_pmf (\_. dflt))) = map_pmf (\f x. if x \ B then f x else dflt) (Pi_pmf B dflt f)" by (simp add: map_pmf_def pair_pmf_def bind_assoc_pmf bind_return_pmf bind_return_pmf') also have "\ = Pi_pmf B dflt f" using assms by (intro Pi_pmf_default_swap) (auto dest: finite_subset) also have "B \ (A - B) = A" using assms by auto finally show ?thesis . qed lemma Pi_pmf_if_set: assumes "finite A" shows "Pi_pmf A dflt (\x. if b x then f x else return_pmf dflt) = Pi_pmf {x\A. b x} dflt f" proof - have "Pi_pmf A dflt (\x. if b x then f x else return_pmf dflt) = Pi_pmf {x\A. b x} dflt (\x. if b x then f x else return_pmf dflt)" using assms by (intro Pi_pmf_subset') auto also have "\ = Pi_pmf {x\A. b x} dflt f" by (intro Pi_pmf_cong) auto finally show ?thesis . qed lemma Pi_pmf_if_set': assumes "finite A" shows "Pi_pmf A dflt (\x. if b x then return_pmf dflt else f x) = Pi_pmf {x\A. \b x} dflt f" proof - have "Pi_pmf A dflt (\x. if b x then return_pmf dflt else f x) = Pi_pmf {x\A. \b x} dflt (\x. if b x then return_pmf dflt else f x)" using assms by (intro Pi_pmf_subset') auto also have "\ = Pi_pmf {x\A. \b x} dflt f" by (intro Pi_pmf_cong) auto finally show ?thesis . qed text \ Lastly, we can delete a single component from a product: \ lemma Pi_pmf_remove: assumes "finite A" shows "Pi_pmf (A - {x}) dflt p = map_pmf (\f. f(x := dflt)) (Pi_pmf A dflt p)" proof - have "Pi_pmf (A - {x}) dflt p = map_pmf (\f xa. if xa \ A - {x} then f xa else dflt) (Pi_pmf A dflt p)" using assms by (intro Pi_pmf_subset) auto also have "\ = map_pmf (\f. f(x := dflt)) (Pi_pmf A dflt p)" using set_Pi_pmf_subset[of A dflt p] assms by (intro map_pmf_cong refl) (auto simp: fun_eq_iff) finally show ?thesis . qed subsection \Additional properties\ lemma nn_integral_prod_Pi_pmf: assumes "finite A" shows "nn_integral (Pi_pmf A dflt p) (\y. \x\A. f x (y x)) = (\x\A. nn_integral (p x) (f x))" using assms proof (induction rule: finite_induct) case (insert x A) have "nn_integral (Pi_pmf (insert x A) dflt p) (\y. \z\insert x A. f z (y z)) = (\\<^sup>+a. \\<^sup>+b. f x a * (\z\A. f z (if z = x then a else b z)) \Pi_pmf A dflt p \p x)" using insert by (auto simp: Pi_pmf_insert case_prod_unfold nn_integral_pair_pmf' cong: if_cong) also have "(\a b. \z\A. f z (if z = x then a else b z)) = (\a b. \z\A. f z (b z))" by (intro ext prod.cong) (use insert.hyps in auto) also have "(\\<^sup>+a. \\<^sup>+b. f x a * (\z\A. f z (b z)) \Pi_pmf A dflt p \p x) = (\\<^sup>+y. f x y \(p x)) * (\\<^sup>+y. (\z\A. f z (y z)) \(Pi_pmf A dflt p))" by (simp add: nn_integral_multc nn_integral_cmult) also have "(\\<^sup>+y. (\z\A. f z (y z)) \(Pi_pmf A dflt p)) = (\x\A. nn_integral (p x) (f x))" by (rule insert.IH) also have "(\\<^sup>+y. f x y \(p x)) * \ = (\x\insert x A. nn_integral (p x) (f x))" using insert.hyps by simp finally show ?case . qed auto lemma integrable_prod_Pi_pmf: fixes f :: "'a \ 'b \ 'c :: {real_normed_field, second_countable_topology, banach}" assumes "finite A" and "\x. x \ A \ integrable (measure_pmf (p x)) (f x)" shows "integrable (measure_pmf (Pi_pmf A dflt p)) (\h. \x\A. f x (h x))" proof (intro integrableI_bounded) have "(\\<^sup>+ x. ennreal (norm (\xa\A. f xa (x xa))) \measure_pmf (Pi_pmf A dflt p)) = (\\<^sup>+ x. (\y\A. ennreal (norm (f y (x y)))) \measure_pmf (Pi_pmf A dflt p))" by (simp flip: prod_norm prod_ennreal) also have "\ = (\x\A. \\<^sup>+ a. ennreal (norm (f x a)) \measure_pmf (p x))" by (intro nn_integral_prod_Pi_pmf) fact also have "(\\<^sup>+a. ennreal (norm (f i a)) \measure_pmf (p i)) \ top" if i: "i \ A" for i using assms(2)[OF i] by (simp add: integrable_iff_bounded) hence "(\x\A. \\<^sup>+ a. ennreal (norm (f x a)) \measure_pmf (p x)) \ top" by (subst ennreal_prod_eq_top) auto finally show "(\\<^sup>+ x. ennreal (norm (\xa\A. f xa (x xa))) \measure_pmf (Pi_pmf A dflt p)) < \" by (simp add: top.not_eq_extremum) qed auto lemma expectation_prod_Pi_pmf: fixes f :: "_ \ _ \ real" assumes "finite A" assumes "\x. x \ A \ integrable (measure_pmf (p x)) (f x)" assumes "\x y. x \ A \ y \ set_pmf (p x) \ f x y \ 0" shows "measure_pmf.expectation (Pi_pmf A dflt p) (\y. \x\A. f x (y x)) = (\x\A. measure_pmf.expectation (p x) (\v. f x v))" proof - have nonneg: "measure_pmf.expectation (p x) (f x) \ 0" if "x \ A" for x using that by (intro Bochner_Integration.integral_nonneg_AE AE_pmfI assms) have nonneg': "0 \ measure_pmf.expectation (Pi_pmf A dflt p) (\y. \x\A. f x (y x))" by (intro Bochner_Integration.integral_nonneg_AE AE_pmfI assms prod_nonneg) (use assms in \auto simp: set_Pi_pmf PiE_dflt_def\) have "ennreal (measure_pmf.expectation (Pi_pmf A dflt p) (\y. \x\A. f x (y x))) = nn_integral (Pi_pmf A dflt p) (\y. ennreal (\x\A. f x (y x)))" using assms by (intro nn_integral_eq_integral [symmetric] assms integrable_prod_Pi_pmf) (auto simp: AE_measure_pmf_iff set_Pi_pmf PiE_dflt_def prod_nonneg) also have "\ = nn_integral (Pi_pmf A dflt p) (\y. (\x\A. ennreal (f x (y x))))" by (intro nn_integral_cong_AE AE_pmfI prod_ennreal [symmetric]) (use assms(1) in \auto simp: set_Pi_pmf PiE_dflt_def intro!: assms(3)\) also have "\ = (\x\A. \\<^sup>+ a. ennreal (f x a) \measure_pmf (p x))" by (rule nn_integral_prod_Pi_pmf) fact+ also have "\ = (\x\A. ennreal (measure_pmf.expectation (p x) (f x)))" by (intro prod.cong nn_integral_eq_integral assms AE_pmfI) auto also have "\ = ennreal (\x\A. measure_pmf.expectation (p x) (f x))" by (intro prod_ennreal nonneg) finally show ?thesis using nonneg nonneg' by (subst (asm) ennreal_inj) (auto intro!: prod_nonneg) qed lemma indep_vars_Pi_pmf: assumes fin: "finite I" shows "prob_space.indep_vars (measure_pmf (Pi_pmf I dflt p)) (\_. count_space UNIV) (\x f. f x) I" proof (cases "I = {}") case True show ?thesis by (subst prob_space.indep_vars_def [OF measure_pmf.prob_space_axioms], subst prob_space.indep_sets_def [OF measure_pmf.prob_space_axioms]) (simp_all add: True) next case [simp]: False show ?thesis proof (subst prob_space.indep_vars_iff_distr_eq_PiM') show "distr (measure_pmf (Pi_pmf I dflt p)) (Pi\<^sub>M I (\i. count_space UNIV)) (\x. restrict x I) = Pi\<^sub>M I (\i. distr (measure_pmf (Pi_pmf I dflt p)) (count_space UNIV) (\f. f i))" proof (rule product_sigma_finite.PiM_eqI, goal_cases) case 1 interpret product_prob_space "\i. distr (measure_pmf (Pi_pmf I dflt p)) (count_space UNIV) (\f. f i)" by (intro product_prob_spaceI prob_space.prob_space_distr measure_pmf.prob_space_axioms) simp_all show ?case by unfold_locales next case 3 have "sets (Pi\<^sub>M I (\i. distr (measure_pmf (Pi_pmf I dflt p)) (count_space UNIV) (\f. f i))) = sets (Pi\<^sub>M I (\_. count_space UNIV))" by (intro sets_PiM_cong) simp_all thus ?case by simp next case (4 A) have "Pi\<^sub>E I A \ sets (Pi\<^sub>M I (\i. count_space UNIV))" using 4 by (intro sets_PiM_I_finite fin) auto hence "emeasure (distr (measure_pmf (Pi_pmf I dflt p)) (Pi\<^sub>M I (\i. count_space UNIV)) (\x. restrict x I)) (Pi\<^sub>E I A) = emeasure (measure_pmf (Pi_pmf I dflt p)) ((\x. restrict x I) -` Pi\<^sub>E I A)" using 4 by (subst emeasure_distr) (auto simp: space_PiM) also have "\ = emeasure (measure_pmf (Pi_pmf I dflt p)) (PiE_dflt I dflt A)" by (intro emeasure_eq_AE AE_pmfI) (auto simp: PiE_dflt_def set_Pi_pmf fin) also have "\ = (\i\I. emeasure (measure_pmf (p i)) (A i))" by (simp add: measure_pmf.emeasure_eq_measure measure_Pi_pmf_PiE_dflt fin prod_ennreal) also have "\ = (\i\I. emeasure (measure_pmf (map_pmf (\f. f i) (Pi_pmf I dflt p))) (A i))" by (intro prod.cong refl, subst Pi_pmf_component) (auto simp: fin) finally show ?case by (simp add: map_pmf_rep_eq) qed fact+ qed (simp_all add: measure_pmf.prob_space_axioms) qed lemma fixes h :: "'a :: comm_monoid_add \ 'b::{banach, second_countable_topology}" assumes fin: "finite I" assumes integrable: "\i. i \ I \ integrable (measure_pmf (D i)) h" shows integrable_sum_Pi_pmf: "integrable (Pi_pmf I dflt D) (\g. \i\I. h (g i))" and expectation_sum_Pi_pmf: "measure_pmf.expectation (Pi_pmf I dflt D) (\g. \i\I. h (g i)) = (\i\I. measure_pmf.expectation (D i) h)" proof - have integrable': "integrable (Pi_pmf I dflt D) (\g. h (g i))" if i: "i \ I" for i proof - have "integrable (D i) h" using i by (rule assms) also have "D i = map_pmf (\g. g i) (Pi_pmf I dflt D)" by (subst Pi_pmf_component) (use fin i in auto) finally show "integrable (measure_pmf (Pi_pmf I dflt D)) (\x. h (x i))" by simp qed thus "integrable (Pi_pmf I dflt D) (\g. \i\I. h (g i))" by (intro Bochner_Integration.integrable_sum) have "measure_pmf.expectation (Pi_pmf I dflt D) (\x. \i\I. h (x i)) = (\i\I. measure_pmf.expectation (map_pmf (\x. x i) (Pi_pmf I dflt D)) h)" using integrable' by (subst Bochner_Integration.integral_sum) auto also have "\ = (\i\I. measure_pmf.expectation (D i) h)" by (intro sum.cong refl, subst Pi_pmf_component) (use fin in auto) finally show "measure_pmf.expectation (Pi_pmf I dflt D) (\g. \i\I. h (g i)) = (\i\I. measure_pmf.expectation (D i) h)" . qed subsection \Applications\ text \ Choosing a subset of a set uniformly at random is equivalent to tossing a fair coin independently for each element and collecting all the elements that came up heads. \ lemma pmf_of_set_Pow_conv_bernoulli: assumes "finite (A :: 'a set)" shows "map_pmf (\b. {x\A. b x}) (Pi_pmf A P (\_. bernoulli_pmf (1/2))) = pmf_of_set (Pow A)" proof - have "Pi_pmf A P (\_. bernoulli_pmf (1/2)) = pmf_of_set (PiE_dflt A P (\x. UNIV))" using assms by (simp add: bernoulli_pmf_half_conv_pmf_of_set Pi_pmf_of_set) also have "map_pmf (\b. {x\A. b x}) \ = pmf_of_set (Pow A)" proof - have "bij_betw (\b. {x \ A. b x}) (PiE_dflt A P (\_. UNIV)) (Pow A)" by (rule bij_betwI[of _ _ _ "\B b. if b \ A then b \ B else P"]) (auto simp add: PiE_dflt_def) then show ?thesis using assms by (intro map_pmf_of_set_bij_betw) auto qed finally show ?thesis by simp qed text \ A binomial distribution can be seen as the number of successes in \n\ independent coin tosses. \ lemma binomial_pmf_altdef': fixes A :: "'a set" assumes "finite A" and "card A = n" and p: "p \ {0..1}" shows "binomial_pmf n p = map_pmf (\f. card {x\A. f x}) (Pi_pmf A dflt (\_. bernoulli_pmf p))" (is "?lhs = ?rhs") proof - from assms have "?lhs = binomial_pmf (card A) p" by simp also have "\ = ?rhs" using assms(1) proof (induction rule: finite_induct) case empty with p show ?case by (simp add: binomial_pmf_0) next case (insert x A) from insert.hyps have "card (insert x A) = Suc (card A)" by simp also have "binomial_pmf \ p = do { b \ bernoulli_pmf p; f \ Pi_pmf A dflt (\_. bernoulli_pmf p); return_pmf ((if b then 1 else 0) + card {y \ A. f y}) }" using p by (simp add: binomial_pmf_Suc insert.IH bind_map_pmf) also have "\ = do { b \ bernoulli_pmf p; f \ Pi_pmf A dflt (\_. bernoulli_pmf p); return_pmf (card {y \ insert x A. (f(x := b)) y}) }" proof (intro bind_pmf_cong refl, goal_cases) case (1 b f) have "(if b then 1 else 0) + card {y\A. f y} = card ((if b then {x} else {}) \ {y\A. f y})" using insert.hyps by auto also have "(if b then {x} else {}) \ {y\A. f y} = {y\insert x A. (f(x := b)) y}" using insert.hyps by auto finally show ?case by simp qed also have "\ = map_pmf (\f. card {y\insert x A. f y}) (Pi_pmf (insert x A) dflt (\_. bernoulli_pmf p))" using insert.hyps by (subst Pi_pmf_insert) (simp_all add: pair_pmf_def map_bind_pmf) finally show ?case . qed finally show ?thesis . qed end diff --git a/src/HOL/Probability/Stream_Space.thy b/src/HOL/Probability/Stream_Space.thy --- a/src/HOL/Probability/Stream_Space.thy +++ b/src/HOL/Probability/Stream_Space.thy @@ -1,668 +1,668 @@ (* Title: HOL/Probability/Stream_Space.thy Author: Johannes Hölzl, TU München *) theory Stream_Space imports Infinite_Product_Measure "HOL-Library.Stream" "HOL-Library.Linear_Temporal_Logic_on_Streams" begin lemma stream_eq_Stream_iff: "s = x ## t \ (shd s = x \ stl s = t)" by (cases s) simp lemma Stream_snth: "(x ## s) !! n = (case n of 0 \ x | Suc n \ s !! n)" by (cases n) simp_all definition to_stream :: "(nat \ 'a) \ 'a stream" where "to_stream X = smap X nats" lemma to_stream_nat_case: "to_stream (case_nat x X) = x ## to_stream X" unfolding to_stream_def by (subst siterate.ctr) (simp add: smap_siterate[symmetric] stream.map_comp comp_def) lemma to_stream_in_streams: "to_stream X \ streams S \ (\n. X n \ S)" by (simp add: to_stream_def streams_iff_snth) definition stream_space :: "'a measure \ 'a stream measure" where "stream_space M = distr (\\<^sub>M i\UNIV. M) (vimage_algebra (streams (space M)) snth (\\<^sub>M i\UNIV. M)) to_stream" lemma space_stream_space: "space (stream_space M) = streams (space M)" by (simp add: stream_space_def) lemma streams_stream_space[intro]: "streams (space M) \ sets (stream_space M)" using sets.top[of "stream_space M"] by (simp add: space_stream_space) lemma stream_space_Stream: "x ## \ \ space (stream_space M) \ x \ space M \ \ \ space (stream_space M)" by (simp add: space_stream_space streams_Stream) lemma stream_space_eq_distr: "stream_space M = distr (\\<^sub>M i\UNIV. M) (stream_space M) to_stream" unfolding stream_space_def by (rule distr_cong) auto lemma sets_stream_space_cong[measurable_cong]: "sets M = sets N \ sets (stream_space M) = sets (stream_space N)" using sets_eq_imp_space_eq[of M N] by (simp add: stream_space_def vimage_algebra_def cong: sets_PiM_cong) lemma measurable_snth_PiM: "(\\ n. \ !! n) \ measurable (stream_space M) (\\<^sub>M i\UNIV. M)" by (auto intro!: measurable_vimage_algebra1 simp: space_PiM streams_iff_sset sset_range image_subset_iff stream_space_def) lemma measurable_snth[measurable]: "(\\. \ !! n) \ measurable (stream_space M) M" using measurable_snth_PiM measurable_component_singleton by (rule measurable_compose) simp lemma measurable_shd[measurable]: "shd \ measurable (stream_space M) M" using measurable_snth[of 0] by simp lemma measurable_stream_space2: assumes f_snth: "\n. (\x. f x !! n) \ measurable N M" shows "f \ measurable N (stream_space M)" unfolding stream_space_def measurable_distr_eq2 proof (rule measurable_vimage_algebra2) show "f \ space N \ streams (space M)" using f_snth[THEN measurable_space] by (auto simp add: streams_iff_sset sset_range) show "(\x. (!!) (f x)) \ measurable N (Pi\<^sub>M UNIV (\i. M))" proof (rule measurable_PiM_single') show "(\x. (!!) (f x)) \ space N \ UNIV \\<^sub>E space M" using f_snth[THEN measurable_space] by auto qed (rule f_snth) qed lemma measurable_stream_coinduct[consumes 1, case_names shd stl, coinduct set: measurable]: assumes "F f" assumes h: "\f. F f \ (\x. shd (f x)) \ measurable N M" assumes t: "\f. F f \ F (\x. stl (f x))" shows "f \ measurable N (stream_space M)" proof (rule measurable_stream_space2) fix n show "(\x. f x !! n) \ measurable N M" using \F f\ by (induction n arbitrary: f) (auto intro: h t) qed lemma measurable_sdrop[measurable]: "sdrop n \ measurable (stream_space M) (stream_space M)" by (rule measurable_stream_space2) (simp add: sdrop_snth) lemma measurable_stl[measurable]: "(\\. stl \) \ measurable (stream_space M) (stream_space M)" by (rule measurable_stream_space2) (simp del: snth.simps add: snth.simps[symmetric]) lemma measurable_to_stream[measurable]: "to_stream \ measurable (\\<^sub>M i\UNIV. M) (stream_space M)" by (rule measurable_stream_space2) (simp add: to_stream_def) lemma measurable_Stream[measurable (raw)]: assumes f[measurable]: "f \ measurable N M" assumes g[measurable]: "g \ measurable N (stream_space M)" shows "(\x. f x ## g x) \ measurable N (stream_space M)" by (rule measurable_stream_space2) (simp add: Stream_snth) lemma measurable_smap[measurable]: assumes X[measurable]: "X \ measurable N M" shows "smap X \ measurable (stream_space N) (stream_space M)" by (rule measurable_stream_space2) simp lemma measurable_stake[measurable]: "stake i \ measurable (stream_space (count_space UNIV)) (count_space (UNIV :: 'a::countable list set))" by (induct i) auto lemma measurable_shift[measurable]: assumes f: "f \ measurable N (stream_space M)" assumes [measurable]: "g \ measurable N (stream_space M)" shows "(\x. stake n (f x) @- g x) \ measurable N (stream_space M)" using f by (induction n arbitrary: f) simp_all lemma measurable_case_stream_replace[measurable (raw)]: "(\x. f x (shd (g x)) (stl (g x))) \ measurable M N \ (\x. case_stream (f x) (g x)) \ measurable M N" unfolding stream.case_eq_if . lemma measurable_ev_at[measurable]: assumes [measurable]: "Measurable.pred (stream_space M) P" shows "Measurable.pred (stream_space M) (ev_at P n)" by (induction n) auto lemma measurable_alw[measurable]: "Measurable.pred (stream_space M) P \ Measurable.pred (stream_space M) (alw P)" unfolding alw_def by (coinduction rule: measurable_gfp_coinduct) (auto simp: inf_continuous_def) lemma measurable_ev[measurable]: "Measurable.pred (stream_space M) P \ Measurable.pred (stream_space M) (ev P)" unfolding ev_def by (coinduction rule: measurable_lfp_coinduct) (auto simp: sup_continuous_def) lemma measurable_until: assumes [measurable]: "Measurable.pred (stream_space M) \" "Measurable.pred (stream_space M) \" shows "Measurable.pred (stream_space M) (\ until \)" unfolding UNTIL_def by (coinduction rule: measurable_gfp_coinduct) (simp_all add: inf_continuous_def fun_eq_iff) lemma measurable_holds [measurable]: "Measurable.pred M P \ Measurable.pred (stream_space M) (holds P)" unfolding holds.simps[abs_def] by (rule measurable_compose[OF measurable_shd]) simp lemma measurable_hld[measurable]: assumes [measurable]: "t \ sets M" shows "Measurable.pred (stream_space M) (HLD t)" unfolding HLD_def by measurable lemma measurable_nxt[measurable (raw)]: "Measurable.pred (stream_space M) P \ Measurable.pred (stream_space M) (nxt P)" unfolding nxt.simps[abs_def] by simp lemma measurable_suntil[measurable]: assumes [measurable]: "Measurable.pred (stream_space M) Q" "Measurable.pred (stream_space M) P" shows "Measurable.pred (stream_space M) (Q suntil P)" unfolding suntil_def by (coinduction rule: measurable_lfp_coinduct) (auto simp: sup_continuous_def) lemma measurable_szip: "(\(\1, \2). szip \1 \2) \ measurable (stream_space M \\<^sub>M stream_space N) (stream_space (M \\<^sub>M N))" proof (rule measurable_stream_space2) fix n have "(\x. (case x of (\1, \2) \ szip \1 \2) !! n) = (\(\1, \2). (\1 !! n, \2 !! n))" by auto also have "\ \ measurable (stream_space M \\<^sub>M stream_space N) (M \\<^sub>M N)" by measurable finally show "(\x. (case x of (\1, \2) \ szip \1 \2) !! n) \ measurable (stream_space M \\<^sub>M stream_space N) (M \\<^sub>M N)" . qed lemma (in prob_space) prob_space_stream_space: "prob_space (stream_space M)" proof - interpret product_prob_space "\_. M" UNIV .. show ?thesis by (subst stream_space_eq_distr) (auto intro!: P.prob_space_distr) qed lemma (in prob_space) nn_integral_stream_space: assumes [measurable]: "f \ borel_measurable (stream_space M)" shows "(\\<^sup>+X. f X \stream_space M) = (\\<^sup>+x. (\\<^sup>+X. f (x ## X) \stream_space M) \M)" proof - interpret S: sequence_space M .. interpret P: pair_sigma_finite M "\\<^sub>M i::nat\UNIV. M" .. have "(\\<^sup>+X. f X \stream_space M) = (\\<^sup>+X. f (to_stream X) \S.S)" by (subst stream_space_eq_distr) (simp add: nn_integral_distr) also have "\ = (\\<^sup>+X. f (to_stream ((\(s, \). case_nat s \) X)) \(M \\<^sub>M S.S))" by (subst S.PiM_iter[symmetric]) (simp add: nn_integral_distr) also have "\ = (\\<^sup>+x. \\<^sup>+X. f (to_stream ((\(s, \). case_nat s \) (x, X))) \S.S \M)" by (subst S.nn_integral_fst) simp_all also have "\ = (\\<^sup>+x. \\<^sup>+X. f (x ## to_stream X) \S.S \M)" by (auto intro!: nn_integral_cong simp: to_stream_nat_case) also have "\ = (\\<^sup>+x. \\<^sup>+X. f (x ## X) \stream_space M \M)" by (subst stream_space_eq_distr) (simp add: nn_integral_distr cong: nn_integral_cong) finally show ?thesis . qed lemma (in prob_space) emeasure_stream_space: assumes X[measurable]: "X \ sets (stream_space M)" shows "emeasure (stream_space M) X = (\\<^sup>+t. emeasure (stream_space M) {x\space (stream_space M). t ## x \ X } \M)" proof - have eq: "\x xs. xs \ space (stream_space M) \ x \ space M \ indicator X (x ## xs) = indicator {xs\space (stream_space M). x ## xs \ X } xs" by (auto split: split_indicator) show ?thesis using nn_integral_stream_space[of "indicator X"] apply (auto intro!: nn_integral_cong) apply (subst nn_integral_cong) apply (rule eq) apply simp_all done qed lemma (in prob_space) prob_stream_space: assumes P[measurable]: "{x\space (stream_space M). P x} \ sets (stream_space M)" shows "\

(x in stream_space M. P x) = (\\<^sup>+t. \

(x in stream_space M. P (t ## x)) \M)" proof - interpret S: prob_space "stream_space M" by (rule prob_space_stream_space) show ?thesis unfolding S.emeasure_eq_measure[symmetric] by (subst emeasure_stream_space) (auto simp: stream_space_Stream intro!: nn_integral_cong) qed lemma (in prob_space) AE_stream_space: assumes [measurable]: "Measurable.pred (stream_space M) P" shows "(AE X in stream_space M. P X) = (AE x in M. AE X in stream_space M. P (x ## X))" proof - interpret stream: prob_space "stream_space M" by (rule prob_space_stream_space) have eq: "\x X. indicator {x. \ P x} (x ## X) = indicator {X. \ P (x ## X)} X" by (auto split: split_indicator) show ?thesis apply (subst AE_iff_nn_integral, simp) apply (subst nn_integral_stream_space, simp) apply (subst eq) apply (subst nn_integral_0_iff_AE, simp) apply (simp add: AE_iff_nn_integral[symmetric]) done qed lemma (in prob_space) AE_stream_all: assumes [measurable]: "Measurable.pred M P" and P: "AE x in M. P x" shows "AE x in stream_space M. stream_all P x" proof - { fix n have "AE x in stream_space M. P (x !! n)" proof (induct n) case 0 with P show ?case by (subst AE_stream_space) (auto elim!: eventually_mono) next case (Suc n) then show ?case by (subst AE_stream_space) auto qed } then show ?thesis unfolding stream_all_def by (simp add: AE_all_countable) qed lemma streams_sets: assumes X[measurable]: "X \ sets M" shows "streams X \ sets (stream_space M)" proof - have "streams X = {x\space (stream_space M). x \ streams X}" using streams_mono[OF _ sets.sets_into_space[OF X]] by (auto simp: space_stream_space) also have "\ = {x\space (stream_space M). gfp (\p x. shd x \ X \ p (stl x)) x}" apply (simp add: set_eq_iff streams_def streamsp_def) apply (intro allI conj_cong refl arg_cong2[where f=gfp] ext) apply (case_tac xa) apply auto done also have "\ \ sets (stream_space M)" apply (intro predE) apply (coinduction rule: measurable_gfp_coinduct) apply (auto simp: inf_continuous_def) done finally show ?thesis . qed lemma sets_stream_space_in_sets: assumes space: "space N = streams (space M)" assumes sets: "\i. (\x. x !! i) \ measurable N M" shows "sets (stream_space M) \ sets N" unfolding stream_space_def sets_distr by (auto intro!: sets_image_in_sets measurable_Sup2 measurable_vimage_algebra2 del: subsetI equalityI simp add: sets_PiM_eq_proj snth_in space sets cong: measurable_cong_sets) lemma sets_stream_space_eq: "sets (stream_space M) = sets (SUP i\UNIV. vimage_algebra (streams (space M)) (\s. s !! i) M)" by (auto intro!: sets_stream_space_in_sets sets_Sup_in_sets sets_image_in_sets measurable_Sup1 snth_in measurable_vimage_algebra1 del: subsetI simp: space_Sup_eq_UN space_stream_space) lemma sets_restrict_stream_space: assumes S[measurable]: "S \ sets M" shows "sets (restrict_space (stream_space M) (streams S)) = sets (stream_space (restrict_space M S))" using S[THEN sets.sets_into_space] apply (subst restrict_space_eq_vimage_algebra) apply (simp add: space_stream_space streams_mono2) apply (subst vimage_algebra_cong[OF refl refl sets_stream_space_eq]) apply (subst sets_stream_space_eq) apply (subst sets_vimage_Sup_eq[where Y="streams (space M)"]) apply simp apply auto [] apply (auto intro: streams_mono) [] apply auto [] apply (simp add: image_image space_restrict_space) apply (simp add: vimage_algebra_cong[OF refl refl restrict_space_eq_vimage_algebra]) apply (subst (1 2) vimage_algebra_vimage_algebra_eq) apply (auto simp: streams_mono snth_in ) done primrec sstart :: "'a set \ 'a list \ 'a stream set" where "sstart S [] = streams S" | [simp del]: "sstart S (x # xs) = (##) x ` sstart S xs" lemma in_sstart[simp]: "s \ sstart S (x # xs) \ shd s = x \ stl s \ sstart S xs" by (cases s) (auto simp: sstart.simps(2)) lemma sstart_in_streams: "xs \ lists S \ sstart S xs \ streams S" by (induction xs) (auto simp: sstart.simps(2)) lemma sstart_eq: "x \ streams S \ x \ sstart S xs = (\i sets (stream_space (count_space UNIV))" proof (induction xs) case (Cons x xs) note Cons[measurable] have "sstart S (x # xs) = {s\space (stream_space (count_space UNIV)). shd s = x \ stl s \ sstart S xs}" by (simp add: set_eq_iff space_stream_space) also have "\ \ sets (stream_space (count_space UNIV))" by measurable finally show ?case . qed (simp add: streams_sets) lemma sigma_sets_singletons: assumes "countable S" shows "sigma_sets S ((\s. {s})`S) = Pow S" proof safe interpret sigma_algebra S "sigma_sets S ((\s. {s})`S)" by (rule sigma_algebra_sigma_sets) auto fix A assume "A \ S" with assms have "(\a\A. {a}) \ sigma_sets S ((\s. {s})`S)" by (intro countable_UN') (auto dest: countable_subset) then show "A \ sigma_sets S ((\s. {s})`S)" by simp qed (auto dest: sigma_sets_into_sp[rotated]) lemma sets_count_space_eq_sigma: "countable S \ sets (count_space S) = sets (sigma S ((\s. {s})`S))" by (subst sets_measure_of) (auto simp: sigma_sets_singletons) lemma sets_stream_space_sstart: assumes S[simp]: "countable S" shows "sets (stream_space (count_space S)) = sets (sigma (streams S) (sstart S`lists S \ {{}}))" proof have [simp]: "sstart S ` lists S \ Pow (streams S)" by (simp add: image_subset_iff sstart_in_streams) let ?S = "sigma (streams S) (sstart S ` lists S \ {{}})" { fix i a assume "a \ S" { fix x have "(x !! i = a \ x \ streams S) \ (\xs\lists S. length xs = i \ x \ sstart S (xs @ [a]))" proof (induction i arbitrary: x) case (Suc i) from this[of "stl x"] show ?case by (simp add: length_Suc_conv Bex_def ex_simps[symmetric] del: ex_simps) (metis stream.collapse streams_Stream) qed (insert \a \ S\, auto intro: streams_stl in_streams) } then have "(\x. x !! i) -` {a} \ streams S = (\xs\{xs\lists S. length xs = i}. sstart S (xs @ [a]))" by (auto simp add: set_eq_iff) also have "\ \ sets ?S" using \a\S\ by (intro sets.countable_UN') (auto intro!: sigma_sets.Basic image_eqI) finally have " (\x. x !! i) -` {a} \ streams S \ sets ?S" . } then show "sets (stream_space (count_space S)) \ sets (sigma (streams S) (sstart S`lists S \ {{}}))" by (intro sets_stream_space_in_sets) (auto simp: measurable_count_space_eq_countable snth_in) have "sigma_sets (space (stream_space (count_space S))) (sstart S`lists S \ {{}}) \ sets (stream_space (count_space S))" proof (safe intro!: sets.sigma_sets_subset) fix xs assume "\x\set xs. x \ S" then have "sstart S xs = {x\space (stream_space (count_space S)). \i \ sets (stream_space (count_space S))" by measurable finally show "sstart S xs \ sets (stream_space (count_space S))" . qed then show "sets (sigma (streams S) (sstart S`lists S \ {{}})) \ sets (stream_space (count_space S))" by (simp add: space_stream_space) qed lemma Int_stable_sstart: "Int_stable (sstart S`lists S \ {{}})" proof - { fix xs ys assume "xs \ lists S" "ys \ lists S" then have "sstart S xs \ sstart S ys \ sstart S ` lists S \ {{}}" proof (induction xs ys rule: list_induct2') case (4 x xs y ys) show ?case proof cases assume "x = y" then have "sstart S (x # xs) \ sstart S (y # ys) = (##) x ` (sstart S xs \ sstart S ys)" by (auto simp: image_iff intro!: stream.collapse[symmetric]) also have "\ \ sstart S ` lists S \ {{}}" using 4 by (auto simp: sstart.simps(2)[symmetric] del: in_listsD) finally show ?case . qed auto qed (simp_all add: sstart_in_streams inf.absorb1 inf.absorb2 image_eqI[where x="[]"]) } then show ?thesis by (auto simp: Int_stable_def) qed lemma stream_space_eq_sstart: assumes S[simp]: "countable S" assumes P: "prob_space M" "prob_space N" assumes ae: "AE x in M. x \ streams S" "AE x in N. x \ streams S" assumes sets_M: "sets M = sets (stream_space (count_space UNIV))" assumes sets_N: "sets N = sets (stream_space (count_space UNIV))" assumes *: "\xs. xs \ [] \ xs \ lists S \ emeasure M (sstart S xs) = emeasure N (sstart S xs)" shows "M = N" proof (rule measure_eqI_restrict_generator[OF Int_stable_sstart]) have [simp]: "sstart S ` lists S \ Pow (streams S)" by (simp add: image_subset_iff sstart_in_streams) interpret M: prob_space M by fact show "sstart S ` lists S \ {{}} \ Pow (streams S)" by (auto dest: sstart_in_streams del: in_listsD) { fix M :: "'a stream measure" assume M: "sets M = sets (stream_space (count_space UNIV))" have "sets (restrict_space M (streams S)) = sigma_sets (streams S) (sstart S ` lists S \ {{}})" by (subst sets_restrict_space_cong[OF M]) (simp add: sets_restrict_stream_space restrict_count_space sets_stream_space_sstart) } from this[OF sets_M] this[OF sets_N] show "sets (restrict_space M (streams S)) = sigma_sets (streams S) (sstart S ` lists S \ {{}})" "sets (restrict_space N (streams S)) = sigma_sets (streams S) (sstart S ` lists S \ {{}})" by auto show "{streams S} \ sstart S ` lists S \ {{}}" "\{streams S} = streams S" "\s. s \ {streams S} \ emeasure M s \ \" using M.emeasure_space_1 space_stream_space[of "count_space S"] sets_eq_imp_space_eq[OF sets_M] by (auto simp add: image_eqI[where x="[]"]) show "sets M = sets N" by (simp add: sets_M sets_N) next fix X assume "X \ sstart S ` lists S \ {{}}" then obtain xs where "X = {} \ (xs \ lists S \ X = sstart S xs)" by auto moreover have "emeasure M (streams S) = 1" using ae by (intro prob_space.emeasure_eq_1_AE[OF P(1)]) (auto simp: sets_M streams_sets) moreover have "emeasure N (streams S) = 1" using ae by (intro prob_space.emeasure_eq_1_AE[OF P(2)]) (auto simp: sets_N streams_sets) ultimately show "emeasure M X = emeasure N X" using P[THEN prob_space.emeasure_space_1] by (cases "xs = []") (auto simp: * space_stream_space del: in_listsD) qed (auto simp: * ae sets_M del: in_listsD intro!: streams_sets) lemma sets_sstart[measurable]: "sstart \ xs \ sets (stream_space (count_space UNIV))" proof (induction xs) case (Cons x xs) note this[measurable] have "sstart \ (x # xs) = {\\space (stream_space (count_space UNIV)). \ \ sstart \ (x # xs)}" by (auto simp: space_stream_space) also have "\ \ sets (stream_space (count_space UNIV))" unfolding in_sstart by measurable finally show ?case . qed (auto intro!: streams_sets) primrec scylinder :: "'a set \ 'a set list \ 'a stream set" where "scylinder S [] = streams S" | "scylinder S (A # As) = {\\streams S. shd \ \ A \ stl \ \ scylinder S As}" lemma scylinder_streams: "scylinder S xs \ streams S" by (induction xs) auto lemma sets_scylinder: "(\x\set xs. x \ sets S) \ scylinder (space S) xs \ sets (stream_space S)" by (induction xs) (auto simp: space_stream_space[symmetric]) lemma stream_space_eq_scylinder: assumes P: "prob_space M" "prob_space N" assumes "Int_stable G" and S: "sets S = sets (sigma (space S) G)" and C: "countable C" "C \ G" "\C = space S" and G: "G \ Pow (space S)" assumes sets_M: "sets M = sets (stream_space S)" assumes sets_N: "sets N = sets (stream_space S)" assumes *: "\xs. xs \ [] \ xs \ lists G \ emeasure M (scylinder (space S) xs) = emeasure N (scylinder (space S) xs)" shows "M = N" proof (rule measure_eqI_generator_eq) interpret M: prob_space M by fact interpret N: prob_space N by fact let ?G = "scylinder (space S) ` lists G" show sc_Pow: "?G \ Pow (streams (space S))" using scylinder_streams by auto have "sets (stream_space S) = sets (sigma (streams (space S)) ?G)" (is "?S = sets ?R") proof (rule antisym) let ?V = "\i. vimage_algebra (streams (space S)) (\s. s !! i) S" show "?S \ sets ?R" unfolding sets_stream_space_eq proof (safe intro!: sets_Sup_in_sets del: subsetI equalityI) fix i :: nat show "space (?V i) = space ?R" - using scylinder_streams by (subst space_measure_of) (auto simp: ) + using scylinder_streams by (subst space_measure_of) auto { fix A assume "A \ G" then have "scylinder (space S) (replicate i (space S) @ [A]) = (\s. s !! i) -` A \ streams (space S)" by (induction i) (auto simp add: streams_shd streams_stl cong: conj_cong) also have "scylinder (space S) (replicate i (space S) @ [A]) = (\xs\{xs\lists C. length xs = i}. scylinder (space S) (xs @ [A]))" apply (induction i) apply auto [] apply (simp add: length_Suc_conv set_eq_iff ex_simps(1,2)[symmetric] cong: conj_cong del: ex_simps(1,2)) apply rule subgoal for i x apply (cases x) apply (subst (2) C(3)[symmetric]) apply (simp del: ex_simps(1,2) add: ex_simps(1,2)[symmetric] ac_simps Bex_def) apply auto done done finally have "(\s. s !! i) -` A \ streams (space S) = (\xs\{xs\lists C. length xs = i}. scylinder (space S) (xs @ [A]))" .. also have "\ \ ?R" using C(2) \A\G\ by (intro sets.countable_UN' countable_Collect countable_lists C) (auto intro!: in_measure_of[OF sc_Pow] imageI) finally have "(\s. s !! i) -` A \ streams (space S) \ ?R" . } then show "sets (?V i) \ ?R" apply (subst vimage_algebra_cong[OF refl refl S]) apply (subst vimage_algebra_sigma[OF G]) apply (simp add: streams_iff_snth) [] apply (subst sigma_le_sets) apply auto done qed have "G \ sets S" unfolding S using G by auto with C(2) show "sets ?R \ ?S" unfolding sigma_le_sets[OF sc_Pow] by (auto intro!: sets_scylinder) qed then show "sets M = sigma_sets (streams (space S)) (scylinder (space S) ` lists G)" "sets N = sigma_sets (streams (space S)) (scylinder (space S) ` lists G)" unfolding sets_M sets_N by (simp_all add: sc_Pow) show "Int_stable ?G" proof (rule Int_stableI_image) fix xs ys assume "xs \ lists G" "ys \ lists G" then show "\zs\lists G. scylinder (space S) xs \ scylinder (space S) ys = scylinder (space S) zs" proof (induction xs arbitrary: ys) case Nil then show ?case by (auto simp add: Int_absorb1 scylinder_streams) next case xs: (Cons x xs) show ?case proof (cases ys) case Nil with xs.hyps show ?thesis by (auto simp add: Int_absorb2 scylinder_streams intro!: bexI[of _ "x#xs"]) next case ys: (Cons y ys') with xs.IH[of ys'] xs.prems obtain zs where "zs \ lists G" and eq: "scylinder (space S) xs \ scylinder (space S) ys' = scylinder (space S) zs" by auto show ?thesis proof (intro bexI[of _ "(x \ y)#zs"]) show "x \ y # zs \ lists G" using \zs\lists G\ \x\G\ \ys\lists G\ ys \Int_stable G\[THEN Int_stableD, of x y] by auto show "scylinder (space S) (x # xs) \ scylinder (space S) ys = scylinder (space S) (x \ y # zs)" by (auto simp add: eq[symmetric] ys) qed qed qed qed show "range (\_::nat. streams (space S)) \ scylinder (space S) ` lists G" "(\i. streams (space S)) = streams (space S)" "emeasure M (streams (space S)) \ \" by (auto intro!: image_eqI[of _ _ "[]"]) fix X assume "X \ scylinder (space S) ` lists G" then obtain xs where xs: "xs \ lists G" and eq: "X = scylinder (space S) xs" by auto then show "emeasure M X = emeasure N X" proof (cases "xs = []") assume "xs = []" then show ?thesis unfolding eq using sets_M[THEN sets_eq_imp_space_eq] sets_N[THEN sets_eq_imp_space_eq] M.emeasure_space_1 N.emeasure_space_1 by (simp add: space_stream_space[symmetric]) next assume "xs \ []" with xs show ?thesis unfolding eq by (intro *) qed qed lemma stream_space_coinduct: fixes R :: "'a stream measure \ 'a stream measure \ bool" assumes "R A B" assumes R: "\A B. R A B \ \K\space (prob_algebra M). \A'\M \\<^sub>M prob_algebra (stream_space M). \B'\M \\<^sub>M prob_algebra (stream_space M). (AE y in K. R (A' y) (B' y) \ A' y = B' y) \ A = do { y \ K; \ \ A' y; return (stream_space M) (y ## \) } \ B = do { y \ K; \ \ B' y; return (stream_space M) (y ## \) }" shows "A = B" proof (rule stream_space_eq_scylinder) let ?step = "\K L. do { y \ K; \ \ L y; return (stream_space M) (y ## \) }" { fix K A A' assume K: "K \ space (prob_algebra M)" and A'[measurable]: "A' \ M \\<^sub>M prob_algebra (stream_space M)" and A_eq: "A = ?step K A'" have ps: "prob_space A" unfolding A_eq by (rule prob_space_bind'[OF K]) measurable have "sets A = sets (stream_space M)" unfolding A_eq by (rule sets_bind'[OF K]) measurable note ps this } note ** = this { fix A B assume "R A B" obtain K A' B' where K: "K \ space (prob_algebra M)" and A': "A' \ M \\<^sub>M prob_algebra (stream_space M)" "A = ?step K A'" and B': "B' \ M \\<^sub>M prob_algebra (stream_space M)" "B = ?step K B'" using R[OF \R A B\] by blast have "prob_space A" "prob_space B" "sets A = sets (stream_space M)" "sets B = sets (stream_space M)" using **[OF K A'] **[OF K B'] by auto } note R_D = this show "prob_space A" "prob_space B" "sets A = sets (stream_space M)" "sets B = sets (stream_space M)" using R_D[OF \R A B\] by auto show "Int_stable (sets M)" "sets M = sets (sigma (space M) (sets M))" "countable {space M}" "{space M} \ sets M" "\{space M} = space M" "sets M \ Pow (space M)" using sets.space_closed[of M] by (auto simp: Int_stable_def) { fix A As L K assume K[measurable]: "K \ space (prob_algebra M)" and A: "A \ sets M" "As \ lists (sets M)" and L[measurable]: "L \ M \\<^sub>M prob_algebra (stream_space M)" from A have [measurable]: "\x\set (A # As). x \ sets M" "\x\set As. x \ sets M" by auto have [simp]: "space K = space M" "sets K = sets M" using K by (auto simp: space_prob_algebra intro!: sets_eq_imp_space_eq) have [simp]: "x \ space M \ sets (L x) = sets (stream_space M)" for x using measurable_space[OF L] by (auto simp: space_prob_algebra) note sets_scylinder[measurable] have *: "indicator (scylinder (space M) (A # As)) (x ## \) = (indicator A x * indicator (scylinder (space M) As) \ :: ennreal)" for \ x using scylinder_streams[of "space M" As] \A \ sets M\[THEN sets.sets_into_space] by (auto split: split_indicator) have "emeasure (?step K L) (scylinder (space M) (A # As)) = (\\<^sup>+y. L y (scylinder (space M) As) * indicator A y \K)" apply (subst emeasure_bind_prob_algebra[OF K]) apply measurable apply (rule nn_integral_cong) apply (subst emeasure_bind_prob_algebra[OF L[THEN measurable_space]]) apply (simp_all add: ac_simps * nn_integral_cmult_indicator del: scylinder.simps) apply measurable done } note emeasure_step = this fix Xs assume "Xs \ lists (sets M)" from this \R A B\ show "emeasure A (scylinder (space M) Xs) = emeasure B (scylinder (space M) Xs)" proof (induction Xs arbitrary: A B) case (Cons X Xs) obtain K A' B' where K: "K \ space (prob_algebra M)" and A'[measurable]: "A' \ M \\<^sub>M prob_algebra (stream_space M)" and A: "A = ?step K A'" and B'[measurable]: "B' \ M \\<^sub>M prob_algebra (stream_space M)" and B: "B = ?step K B'" and AE_R: "AE x in K. R (A' x) (B' x) \ A' x = B' x" using R[OF \R A B\] by blast show ?case unfolding A B emeasure_step[OF K Cons.hyps A'] emeasure_step[OF K Cons.hyps B'] apply (rule nn_integral_cong_AE) using AE_R by eventually_elim (auto simp add: Cons.IH) next case Nil note R_D[OF this] from this(1,2)[THEN prob_space.emeasure_space_1] this(3,4)[THEN sets_eq_imp_space_eq] show ?case by (simp add: space_stream_space) qed qed end diff --git a/src/HOL/Quotient.thy b/src/HOL/Quotient.thy --- a/src/HOL/Quotient.thy +++ b/src/HOL/Quotient.thy @@ -1,883 +1,883 @@ (* Title: HOL/Quotient.thy Author: Cezary Kaliszyk and Christian Urban *) section \Definition of Quotient Types\ theory Quotient imports Lifting keywords "print_quotmapsQ3" "print_quotientsQ3" "print_quotconsts" :: diag and "quotient_type" :: thy_goal_defn and "/" and "quotient_definition" :: thy_goal_defn and "copy_bnf" :: thy_defn and "lift_bnf" :: thy_goal_defn begin text \ Basic definition for equivalence relations that are represented by predicates. \ text \Composition of Relations\ abbreviation rel_conj :: "('a \ 'b \ bool) \ ('b \ 'a \ bool) \ 'a \ 'b \ bool" (infixr "OOO" 75) where "r1 OOO r2 \ r1 OO r2 OO r1" lemma eq_comp_r: shows "((=) OOO R) = R" by (auto simp add: fun_eq_iff) context includes lifting_syntax begin subsection \Quotient Predicate\ definition "Quotient3 R Abs Rep \ (\a. Abs (Rep a) = a) \ (\a. R (Rep a) (Rep a)) \ (\r s. R r s \ R r r \ R s s \ Abs r = Abs s)" lemma Quotient3I: assumes "\a. Abs (Rep a) = a" and "\a. R (Rep a) (Rep a)" and "\r s. R r s \ R r r \ R s s \ Abs r = Abs s" shows "Quotient3 R Abs Rep" using assms unfolding Quotient3_def by blast context fixes R Abs Rep assumes a: "Quotient3 R Abs Rep" begin lemma Quotient3_abs_rep: "Abs (Rep a) = a" using a unfolding Quotient3_def by simp lemma Quotient3_rep_reflp: "R (Rep a) (Rep a)" using a unfolding Quotient3_def by blast lemma Quotient3_rel: "R r r \ R s s \ Abs r = Abs s \ R r s" \ \orientation does not loop on rewriting\ using a unfolding Quotient3_def by blast lemma Quotient3_refl1: "R r s \ R r r" using a unfolding Quotient3_def by fast lemma Quotient3_refl2: "R r s \ R s s" using a unfolding Quotient3_def by fast lemma Quotient3_rel_rep: "R (Rep a) (Rep b) \ a = b" using a unfolding Quotient3_def by metis lemma Quotient3_rep_abs: "R r r \ R (Rep (Abs r)) r" using a unfolding Quotient3_def by blast lemma Quotient3_rel_abs: "R r s \ Abs r = Abs s" using a unfolding Quotient3_def by blast lemma Quotient3_symp: "symp R" using a unfolding Quotient3_def using sympI by metis lemma Quotient3_transp: "transp R" using a unfolding Quotient3_def using transpI by (metis (full_types)) lemma Quotient3_part_equivp: "part_equivp R" by (metis Quotient3_rep_reflp Quotient3_symp Quotient3_transp part_equivpI) lemma abs_o_rep: "Abs \ Rep = id" unfolding fun_eq_iff by (simp add: Quotient3_abs_rep) lemma equals_rsp: assumes b: "R xa xb" "R ya yb" shows "R xa ya = R xb yb" using b Quotient3_symp Quotient3_transp by (blast elim: sympE transpE) lemma rep_abs_rsp: assumes b: "R x1 x2" shows "R x1 (Rep (Abs x2))" using b Quotient3_rel Quotient3_abs_rep Quotient3_rep_reflp by metis lemma rep_abs_rsp_left: assumes b: "R x1 x2" shows "R (Rep (Abs x1)) x2" using b Quotient3_rel Quotient3_abs_rep Quotient3_rep_reflp by metis end lemma identity_quotient3: "Quotient3 (=) id id" unfolding Quotient3_def id_def by blast lemma fun_quotient3: assumes q1: "Quotient3 R1 abs1 rep1" and q2: "Quotient3 R2 abs2 rep2" shows "Quotient3 (R1 ===> R2) (rep1 ---> abs2) (abs1 ---> rep2)" proof - have "(rep1 ---> abs2) ((abs1 ---> rep2) a) = a" for a using q1 q2 by (simp add: Quotient3_def fun_eq_iff) moreover have "(R1 ===> R2) ((abs1 ---> rep2) a) ((abs1 ---> rep2) a)" for a by (rule rel_funI) (insert q1 q2 Quotient3_rel_abs [of R1 abs1 rep1] Quotient3_rel_rep [of R2 abs2 rep2], simp (no_asm) add: Quotient3_def, simp) moreover have "(R1 ===> R2) r s = ((R1 ===> R2) r r \ (R1 ===> R2) s s \ (rep1 ---> abs2) r = (rep1 ---> abs2) s)" for r s proof - have "(R1 ===> R2) r s \ (R1 ===> R2) r r" unfolding rel_fun_def using Quotient3_part_equivp[OF q1] Quotient3_part_equivp[OF q2] by (metis (full_types) part_equivp_def) moreover have "(R1 ===> R2) r s \ (R1 ===> R2) s s" unfolding rel_fun_def using Quotient3_part_equivp[OF q1] Quotient3_part_equivp[OF q2] by (metis (full_types) part_equivp_def) moreover have "(R1 ===> R2) r s \ (rep1 ---> abs2) r = (rep1 ---> abs2) s" by (auto simp add: rel_fun_def fun_eq_iff) (use q1 q2 in \unfold Quotient3_def, metis\) moreover have "((R1 ===> R2) r r \ (R1 ===> R2) s s \ (rep1 ---> abs2) r = (rep1 ---> abs2) s) \ (R1 ===> R2) r s" by (auto simp add: rel_fun_def fun_eq_iff) (use q1 q2 in \unfold Quotient3_def, metis map_fun_apply\) ultimately show ?thesis by blast qed ultimately show ?thesis by (intro Quotient3I) (assumption+) qed lemma lambda_prs: assumes q1: "Quotient3 R1 Abs1 Rep1" and q2: "Quotient3 R2 Abs2 Rep2" shows "(Rep1 ---> Abs2) (\x. Rep2 (f (Abs1 x))) = (\x. f x)" unfolding fun_eq_iff using Quotient3_abs_rep[OF q1] Quotient3_abs_rep[OF q2] by simp lemma lambda_prs1: assumes q1: "Quotient3 R1 Abs1 Rep1" and q2: "Quotient3 R2 Abs2 Rep2" shows "(Rep1 ---> Abs2) (\x. (Abs1 ---> Rep2) f x) = (\x. f x)" unfolding fun_eq_iff using Quotient3_abs_rep[OF q1] Quotient3_abs_rep[OF q2] by simp text\ In the following theorem R1 can be instantiated with anything, but we know some of the types of the Rep and Abs functions; so by solving Quotient assumptions we can get a unique R1 that will be provable; which is why we need to use \apply_rsp\ and not the primed version\ lemma apply_rspQ3: fixes f g::"'a \ 'c" assumes q: "Quotient3 R1 Abs1 Rep1" and a: "(R1 ===> R2) f g" "R1 x y" shows "R2 (f x) (g y)" using a by (auto elim: rel_funE) lemma apply_rspQ3'': assumes "Quotient3 R Abs Rep" and "(R ===> S) f f" shows "S (f (Rep x)) (f (Rep x))" proof - from assms(1) have "R (Rep x) (Rep x)" by (rule Quotient3_rep_reflp) then show ?thesis using assms(2) by (auto intro: apply_rsp') qed subsection \lemmas for regularisation of ball and bex\ lemma ball_reg_eqv: fixes P :: "'a \ bool" assumes a: "equivp R" shows "Ball (Respects R) P = (All P)" using a unfolding equivp_def by (auto simp add: in_respects) lemma bex_reg_eqv: fixes P :: "'a \ bool" assumes a: "equivp R" shows "Bex (Respects R) P = (Ex P)" using a unfolding equivp_def by (auto simp add: in_respects) lemma ball_reg_right: assumes a: "\x. x \ R \ P x \ Q x" shows "All P \ Ball R Q" using a by fast lemma bex_reg_left: assumes a: "\x. x \ R \ Q x \ P x" shows "Bex R Q \ Ex P" using a by fast lemma ball_reg_left: assumes a: "equivp R" shows "(\x. (Q x \ P x)) \ Ball (Respects R) Q \ All P" using a by (metis equivp_reflp in_respects) lemma bex_reg_right: assumes a: "equivp R" shows "(\x. (Q x \ P x)) \ Ex Q \ Bex (Respects R) P" using a by (metis equivp_reflp in_respects) lemma ball_reg_eqv_range: fixes P::"'a \ bool" and x::"'a" assumes a: "equivp R2" shows "(Ball (Respects (R1 ===> R2)) (\f. P (f x)) = All (\f. P (f x)))" proof (intro allI iffI) fix f assume "\f \ Respects (R1 ===> R2). P (f x)" moreover have "(\y. f x) \ Respects (R1 ===> R2)" using a equivp_reflp_symp_transp[of "R2"] by(auto simp add: in_respects rel_fun_def elim: equivpE reflpE) ultimately show "P (f x)" by auto qed auto lemma bex_reg_eqv_range: assumes a: "equivp R2" shows "(Bex (Respects (R1 ===> R2)) (\f. P (f x)) = Ex (\f. P (f x)))" proof - { fix f assume "P (f x)" have "(\y. f x) \ Respects (R1 ===> R2)" using a equivp_reflp_symp_transp[of "R2"] by (auto simp add: Respects_def in_respects rel_fun_def elim: equivpE reflpE) } then show ?thesis by auto qed (* Next four lemmas are unused *) lemma all_reg: assumes a: "\x :: 'a. (P x \ Q x)" and b: "All P" shows "All Q" using a b by fast lemma ex_reg: assumes a: "\x :: 'a. (P x \ Q x)" and b: "Ex P" shows "Ex Q" using a b by fast lemma ball_reg: assumes a: "\x :: 'a. (x \ R \ P x \ Q x)" and b: "Ball R P" shows "Ball R Q" using a b by fast lemma bex_reg: assumes a: "\x :: 'a. (x \ R \ P x \ Q x)" and b: "Bex R P" shows "Bex R Q" using a b by fast lemma ball_all_comm: assumes "\y. (\x\P. A x y) \ (\x. B x y)" shows "(\x\P. \y. A x y) \ (\x. \y. B x y)" using assms by auto lemma bex_ex_comm: assumes "(\y. \x. A x y) \ (\y. \x\P. B x y)" shows "(\x. \y. A x y) \ (\x\P. \y. B x y)" using assms by auto subsection \Bounded abstraction\ definition Babs :: "'a set \ ('a \ 'b) \ 'a \ 'b" where "x \ p \ Babs p m x = m x" lemma babs_rsp: assumes q: "Quotient3 R1 Abs1 Rep1" and a: "(R1 ===> R2) f g" shows "(R1 ===> R2) (Babs (Respects R1) f) (Babs (Respects R1) g)" proof (clarsimp simp add: Babs_def in_respects rel_fun_def) fix x y assume "R1 x y" then have "x \ Respects R1 \ y \ Respects R1" unfolding in_respects rel_fun_def using Quotient3_rel[OF q]by metis then show "R2 (Babs (Respects R1) f x) (Babs (Respects R1) g y)" using \R1 x y\ a by (simp add: Babs_def rel_fun_def) qed lemma babs_prs: assumes q1: "Quotient3 R1 Abs1 Rep1" and q2: "Quotient3 R2 Abs2 Rep2" shows "((Rep1 ---> Abs2) (Babs (Respects R1) ((Abs1 ---> Rep2) f))) = f" proof - { fix x have "Rep1 x \ Respects R1" by (simp add: in_respects Quotient3_rel_rep[OF q1]) then have "Abs2 (Babs (Respects R1) ((Abs1 ---> Rep2) f) (Rep1 x)) = f x" by (simp add: Babs_def Quotient3_abs_rep[OF q1] Quotient3_abs_rep[OF q2]) } then show ?thesis by force qed lemma babs_simp: assumes q: "Quotient3 R1 Abs Rep" shows "((R1 ===> R2) (Babs (Respects R1) f) (Babs (Respects R1) g)) = ((R1 ===> R2) f g)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs unfolding rel_fun_def by (metis Babs_def in_respects Quotient3_rel[OF q]) qed (simp add: babs_rsp[OF q]) text \If a user proves that a particular functional relation is an equivalence, this may be useful in regularising\ lemma babs_reg_eqv: shows "equivp R \ Babs (Respects R) P = P" by (simp add: fun_eq_iff Babs_def in_respects equivp_reflp) (* 3 lemmas needed for proving repabs_inj *) lemma ball_rsp: assumes a: "(R ===> (=)) f g" shows "Ball (Respects R) f = Ball (Respects R) g" using a by (auto simp add: Ball_def in_respects elim: rel_funE) lemma bex_rsp: assumes a: "(R ===> (=)) f g" shows "(Bex (Respects R) f = Bex (Respects R) g)" using a by (auto simp add: Bex_def in_respects elim: rel_funE) lemma bex1_rsp: assumes a: "(R ===> (=)) f g" shows "Ex1 (\x. x \ Respects R \ f x) = Ex1 (\x. x \ Respects R \ g x)" using a by (auto elim: rel_funE simp add: Ex1_def in_respects) text \Two lemmas needed for cleaning of quantifiers\ lemma all_prs: assumes a: "Quotient3 R absf repf" shows "Ball (Respects R) ((absf ---> id) f) = All f" using a unfolding Quotient3_def Ball_def in_respects id_apply comp_def map_fun_def by metis lemma ex_prs: assumes a: "Quotient3 R absf repf" shows "Bex (Respects R) ((absf ---> id) f) = Ex f" using a unfolding Quotient3_def Bex_def in_respects id_apply comp_def map_fun_def by metis subsection \\Bex1_rel\ quantifier\ definition Bex1_rel :: "('a \ 'a \ bool) \ ('a \ bool) \ bool" where "Bex1_rel R P \ (\x \ Respects R. P x) \ (\x \ Respects R. \y \ Respects R. ((P x \ P y) \ (R x y)))" lemma bex1_rel_aux: "\\xa ya. R xa ya \ x xa = y ya; Bex1_rel R x\ \ Bex1_rel R y" unfolding Bex1_rel_def by (metis in_respects) lemma bex1_rel_aux2: "\\xa ya. R xa ya \ x xa = y ya; Bex1_rel R y\ \ Bex1_rel R x" unfolding Bex1_rel_def by (metis in_respects) lemma bex1_rel_rsp: assumes a: "Quotient3 R absf repf" shows "((R ===> (=)) ===> (=)) (Bex1_rel R) (Bex1_rel R)" unfolding rel_fun_def by (metis bex1_rel_aux bex1_rel_aux2) lemma ex1_prs: assumes "Quotient3 R absf repf" shows "((absf ---> id) ---> id) (Bex1_rel R) f = Ex1 f" (is "?lhs = ?rhs") using assms apply (auto simp add: Bex1_rel_def Respects_def) by (metis (full_types) Quotient3_def)+ lemma bex1_bexeq_reg: shows "(\!x\Respects R. P x) \ (Bex1_rel R (\x. P x))" by (auto simp add: Ex1_def Bex1_rel_def Bex_def Ball_def in_respects) lemma bex1_bexeq_reg_eqv: assumes a: "equivp R" shows "(\!x. P x) \ Bex1_rel R P" using equivp_reflp[OF a] by (metis (full_types) Bex1_rel_def in_respects) subsection \Various respects and preserve lemmas\ lemma quot_rel_rsp: assumes a: "Quotient3 R Abs Rep" shows "(R ===> R ===> (=)) R R" apply(rule rel_funI)+ by (meson assms equals_rsp) lemma o_prs: assumes q1: "Quotient3 R1 Abs1 Rep1" and q2: "Quotient3 R2 Abs2 Rep2" and q3: "Quotient3 R3 Abs3 Rep3" shows "((Abs2 ---> Rep3) ---> (Abs1 ---> Rep2) ---> (Rep1 ---> Abs3)) (\) = (\)" and "(id ---> (Abs1 ---> id) ---> Rep1 ---> id) (\) = (\)" using Quotient3_abs_rep[OF q1] Quotient3_abs_rep[OF q2] Quotient3_abs_rep[OF q3] by (simp_all add: fun_eq_iff) lemma o_rsp: "((R2 ===> R3) ===> (R1 ===> R2) ===> (R1 ===> R3)) (\) (\)" "((=) ===> (R1 ===> (=)) ===> R1 ===> (=)) (\) (\)" by (force elim: rel_funE)+ lemma cond_prs: assumes a: "Quotient3 R absf repf" shows "absf (if a then repf b else repf c) = (if a then b else c)" using a unfolding Quotient3_def by auto lemma if_prs: assumes q: "Quotient3 R Abs Rep" shows "(id ---> Rep ---> Rep ---> Abs) If = If" using Quotient3_abs_rep[OF q] by (auto simp add: fun_eq_iff) lemma if_rsp: assumes q: "Quotient3 R Abs Rep" shows "((=) ===> R ===> R ===> R) If If" by force lemma let_prs: assumes q1: "Quotient3 R1 Abs1 Rep1" and q2: "Quotient3 R2 Abs2 Rep2" shows "(Rep2 ---> (Abs2 ---> Rep1) ---> Abs1) Let = Let" using Quotient3_abs_rep[OF q1] Quotient3_abs_rep[OF q2] by (auto simp add: fun_eq_iff) lemma let_rsp: shows "(R1 ===> (R1 ===> R2) ===> R2) Let Let" by (force elim: rel_funE) lemma id_rsp: shows "(R ===> R) id id" by auto lemma id_prs: assumes a: "Quotient3 R Abs Rep" shows "(Rep ---> Abs) id = id" by (simp add: fun_eq_iff Quotient3_abs_rep [OF a]) end locale quot_type = fixes R :: "'a \ 'a \ bool" and Abs :: "'a set \ 'b" and Rep :: "'b \ 'a set" assumes equivp: "part_equivp R" and rep_prop: "\y. \x. R x x \ Rep y = Collect (R x)" and rep_inverse: "\x. Abs (Rep x) = x" and abs_inverse: "\c. (\x. ((R x x) \ (c = Collect (R x)))) \ (Rep (Abs c)) = c" and rep_inject: "\x y. (Rep x = Rep y) = (x = y)" begin definition abs :: "'a \ 'b" where "abs x = Abs (Collect (R x))" definition rep :: "'b \ 'a" where "rep a = (SOME x. x \ Rep a)" lemma some_collect: assumes "R r r" shows "R (SOME x. x \ Collect (R r)) = R r" apply simp by (metis assms exE_some equivp[simplified part_equivp_def]) lemma Quotient: shows "Quotient3 R abs rep" unfolding Quotient3_def abs_def rep_def proof (intro conjI allI) fix a r s show x: "R (SOME x. x \ Rep a) (SOME x. x \ Rep a)" proof - obtain x where r: "R x x" and rep: "Rep a = Collect (R x)" using rep_prop[of a] by auto have "R (SOME x. x \ Rep a) x" using r rep some_collect by metis then have "R x (SOME x. x \ Rep a)" using part_equivp_symp[OF equivp] by fast then show "R (SOME x. x \ Rep a) (SOME x. x \ Rep a)" using part_equivp_transp[OF equivp] by (metis \R (SOME x. x \ Rep a) x\) qed have "Collect (R (SOME x. x \ Rep a)) = (Rep a)" by (metis some_collect rep_prop) then show "Abs (Collect (R (SOME x. x \ Rep a))) = a" using rep_inverse by auto have "R r r \ R s s \ Abs (Collect (R r)) = Abs (Collect (R s)) \ R r = R s" proof - assume "R r r" and "R s s" then have "Abs (Collect (R r)) = Abs (Collect (R s)) \ Collect (R r) = Collect (R s)" by (metis abs_inverse) also have "Collect (R r) = Collect (R s) \ (\A x. x \ A) (Collect (R r)) = (\A x. x \ A) (Collect (R s))" by rule simp_all finally show "Abs (Collect (R r)) = Abs (Collect (R s)) \ R r = R s" by simp qed then show "R r s \ R r r \ R s s \ (Abs (Collect (R r)) = Abs (Collect (R s)))" using equivp[simplified part_equivp_def] by metis qed end subsection \Quotient composition\ lemma OOO_quotient3: fixes R1 :: "'a \ 'a \ bool" fixes Abs1 :: "'a \ 'b" and Rep1 :: "'b \ 'a" fixes Abs2 :: "'b \ 'c" and Rep2 :: "'c \ 'b" fixes R2' :: "'a \ 'a \ bool" fixes R2 :: "'b \ 'b \ bool" assumes R1: "Quotient3 R1 Abs1 Rep1" assumes R2: "Quotient3 R2 Abs2 Rep2" assumes Abs1: "\x y. R2' x y \ R1 x x \ R1 y y \ R2 (Abs1 x) (Abs1 y)" assumes Rep1: "\x y. R2 x y \ R2' (Rep1 x) (Rep1 y)" shows "Quotient3 (R1 OO R2' OO R1) (Abs2 \ Abs1) (Rep1 \ Rep2)" proof - have *: "(R1 OOO R2') r r \ (R1 OOO R2') s s \ (Abs2 \ Abs1) r = (Abs2 \ Abs1) s \ (R1 OOO R2') r s" for r s proof (intro iffI conjI; clarify) show "(R1 OOO R2') r s" if r: "R1 r a" "R2' a b" "R1 b r" and eq: "(Abs2 \ Abs1) r = (Abs2 \ Abs1) s" and s: "R1 s c" "R2' c d" "R1 d s" for a b c d proof - have "R1 r (Rep1 (Abs1 r))" using r Quotient3_refl1 R1 rep_abs_rsp by fastforce moreover have "R2' (Rep1 (Abs1 r)) (Rep1 (Abs1 s))" using that - apply (simp add: ) + apply simp apply (metis (full_types) Rep1 Abs1 Quotient3_rel R2 Quotient3_refl1 [OF R1] Quotient3_refl2 [OF R1] Quotient3_rel_abs [OF R1]) done moreover have "R1 (Rep1 (Abs1 s)) s" by (metis s Quotient3_rel R1 rep_abs_rsp_left) ultimately show ?thesis by (metis relcomppI) qed next fix x y assume xy: "R1 r x" "R2' x y" "R1 y s" then have "R2 (Abs1 x) (Abs1 y)" by (iprover dest: Abs1 elim: Quotient3_refl1 [OF R1] Quotient3_refl2 [OF R1]) then have "R2' (Rep1 (Abs1 x)) (Rep1 (Abs1 x))" "R2' (Rep1 (Abs1 y)) (Rep1 (Abs1 y))" by (simp_all add: Quotient3_refl1 [OF R2] Quotient3_refl2 [OF R2] Rep1) with \R1 r x\ \R1 y s\ show "(R1 OOO R2') r r" "(R1 OOO R2') s s" by (metis (full_types) Quotient3_def R1 relcompp.relcompI)+ show "(Abs2 \ Abs1) r = (Abs2 \ Abs1) s" using xy by simp (metis (full_types) Abs1 Quotient3_rel R1 R2) qed show ?thesis apply (rule Quotient3I) using * apply (simp_all add: o_def Quotient3_abs_rep [OF R2] Quotient3_abs_rep [OF R1]) apply (metis Quotient3_rep_reflp R1 R2 Rep1 relcompp.relcompI) done qed lemma OOO_eq_quotient3: fixes R1 :: "'a \ 'a \ bool" fixes Abs1 :: "'a \ 'b" and Rep1 :: "'b \ 'a" fixes Abs2 :: "'b \ 'c" and Rep2 :: "'c \ 'b" assumes R1: "Quotient3 R1 Abs1 Rep1" assumes R2: "Quotient3 (=) Abs2 Rep2" shows "Quotient3 (R1 OOO (=)) (Abs2 \ Abs1) (Rep1 \ Rep2)" using assms by (rule OOO_quotient3) auto subsection \Quotient3 to Quotient\ lemma Quotient3_to_Quotient: assumes "Quotient3 R Abs Rep" and "T \ \x y. R x x \ Abs x = y" shows "Quotient R Abs Rep T" using assms unfolding Quotient3_def by (intro QuotientI) blast+ lemma Quotient3_to_Quotient_equivp: assumes q: "Quotient3 R Abs Rep" and T_def: "T \ \x y. Abs x = y" and eR: "equivp R" shows "Quotient R Abs Rep T" proof (intro QuotientI) fix a show "Abs (Rep a) = a" using q by(rule Quotient3_abs_rep) next fix a show "R (Rep a) (Rep a)" using q by(rule Quotient3_rep_reflp) next fix r s show "R r s = (R r r \ R s s \ Abs r = Abs s)" using q by(rule Quotient3_rel[symmetric]) next show "T = (\x y. R x x \ Abs x = y)" using T_def equivp_reflp[OF eR] by simp qed subsection \ML setup\ text \Auxiliary data for the quotient package\ named_theorems quot_equiv "equivalence relation theorems" and quot_respect "respectfulness theorems" and quot_preserve "preservation theorems" and id_simps "identity simp rules for maps" and quot_thm "quotient theorems" ML_file \Tools/Quotient/quotient_info.ML\ declare [[mapQ3 "fun" = (rel_fun, fun_quotient3)]] lemmas [quot_thm] = fun_quotient3 lemmas [quot_respect] = quot_rel_rsp if_rsp o_rsp let_rsp id_rsp lemmas [quot_preserve] = if_prs o_prs let_prs id_prs lemmas [quot_equiv] = identity_equivp text \Lemmas about simplifying id's.\ lemmas [id_simps] = id_def[symmetric] map_fun_id id_apply id_o o_id eq_comp_r vimage_id text \Translation functions for the lifting process.\ ML_file \Tools/Quotient/quotient_term.ML\ text \Definitions of the quotient types.\ ML_file \Tools/Quotient/quotient_type.ML\ text \Definitions for quotient constants.\ ML_file \Tools/Quotient/quotient_def.ML\ text \ An auxiliary constant for recording some information about the lifted theorem in a tactic. \ definition Quot_True :: "'a \ bool" where "Quot_True x \ True" lemma shows QT_all: "Quot_True (All P) \ Quot_True P" and QT_ex: "Quot_True (Ex P) \ Quot_True P" and QT_ex1: "Quot_True (Ex1 P) \ Quot_True P" and QT_lam: "Quot_True (\x. P x) \ (\x. Quot_True (P x))" and QT_ext: "(\x. Quot_True (a x) \ f x = g x) \ (Quot_True a \ f = g)" by (simp_all add: Quot_True_def ext) lemma QT_imp: "Quot_True a \ Quot_True b" by (simp add: Quot_True_def) context includes lifting_syntax begin text \Tactics for proving the lifted theorems\ ML_file \Tools/Quotient/quotient_tacs.ML\ end subsection \Methods / Interface\ method_setup lifting = \Attrib.thms >> (fn thms => fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.lift_tac ctxt [] thms))\ \lift theorems to quotient types\ method_setup lifting_setup = \Attrib.thm >> (fn thm => fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.lift_procedure_tac ctxt [] thm))\ \set up the three goals for the quotient lifting procedure\ method_setup descending = \Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.descend_tac ctxt []))\ \decend theorems to the raw level\ method_setup descending_setup = \Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.descend_procedure_tac ctxt []))\ \set up the three goals for the decending theorems\ method_setup partiality_descending = \Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.partiality_descend_tac ctxt []))\ \decend theorems to the raw level\ method_setup partiality_descending_setup = \Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.partiality_descend_procedure_tac ctxt []))\ \set up the three goals for the decending theorems\ method_setup regularize = \Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.regularize_tac ctxt))\ \prove the regularization goals from the quotient lifting procedure\ method_setup injection = \Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.all_injection_tac ctxt))\ \prove the rep/abs injection goals from the quotient lifting procedure\ method_setup cleaning = \Scan.succeed (fn ctxt => SIMPLE_METHOD' (Quotient_Tacs.clean_tac ctxt))\ \prove the cleaning goals from the quotient lifting procedure\ attribute_setup quot_lifted = \Scan.succeed Quotient_Tacs.lifted_attrib\ \lift theorems to quotient types\ no_notation rel_conj (infixr "OOO" 75) section \Lifting of BNFs\ lemma sum_insert_Inl_unit: "x \ A \ (\y. x = Inr y \ Inr y \ B) \ x \ insert (Inl ()) B" by (cases x) (simp_all) lemma lift_sum_unit_vimage_commute: "insert (Inl ()) (Inr ` f -` A) = map_sum id f -` insert (Inl ()) (Inr ` A)" by (auto simp: map_sum_def split: sum.splits) lemma insert_Inl_int_map_sum_unit: "insert (Inl ()) A \ range (map_sum id f) \ {}" by (auto simp: map_sum_def split: sum.splits) lemma image_map_sum_unit_subset: "A \ insert (Inl ()) (Inr ` B) \ map_sum id f ` A \ insert (Inl ()) (Inr ` f ` B)" by auto lemma subset_lift_sum_unitD: "A \ insert (Inl ()) (Inr ` B) \ Inr x \ A \ x \ B" unfolding insert_def by auto lemma UNIV_sum_unit_conv: "insert (Inl ()) (range Inr) = UNIV" unfolding UNIV_sum UNIV_unit image_insert image_empty Un_insert_left sup_bot.left_neutral.. lemma subset_vimage_image_subset: "A \ f -` B \ f ` A \ B" by auto lemma relcompp_mem_Grp_neq_bot: "A \ range f \ {} \ (\x y. x \ A \ y \ A) OO (Grp UNIV f)\\ \ bot" unfolding Grp_def relcompp_apply fun_eq_iff by blast lemma comp_projr_Inr: "projr \ Inr = id" by auto lemma in_rel_sum_in_image_projr: "B \ {(x,y). rel_sum ((=) :: unit \ unit \ bool) A x y} \ Inr ` C = fst ` B \ snd ` B = Inr ` D \ map_prod projr projr ` B \ {(x,y). A x y}" by (force simp: projr_def image_iff dest!: spec[of _ "Inl ()"] split: sum.splits) lemma subset_rel_sumI: "B \ {(x,y). A x y} \ rel_sum ((=) :: unit => unit => bool) A (if x \ B then Inr (fst x) else Inl ()) (if x \ B then Inr (snd x) else Inl ())" by auto lemma relcompp_eq_Grp_neq_bot: "(=) OO (Grp UNIV f)\\ \ bot" unfolding Grp_def relcompp_apply fun_eq_iff by blast lemma rel_fun_rel_OO1: "(rel_fun Q (rel_fun R (=))) A B \ conversep Q OO A OO R \ B" by (auto simp: rel_fun_def) lemma rel_fun_rel_OO2: "(rel_fun Q (rel_fun R (=))) A B \ Q OO B OO conversep R \ A" by (auto simp: rel_fun_def) lemma rel_sum_eq2_nonempty: "rel_sum (=) A OO rel_sum (=) B \ bot" by (auto simp: fun_eq_iff relcompp_apply intro!: exI[of _ "Inl _"]) lemma rel_sum_eq3_nonempty: "rel_sum (=) A OO (rel_sum (=) B OO rel_sum (=) C) \ bot" by (auto simp: fun_eq_iff relcompp_apply intro!: exI[of _ "Inl _"]) lemma hypsubst: "A = B \ x \ B \ (x \ A \ P) \ P" by simp lemma Quotient_crel_quotient: "Quotient R Abs Rep T \ equivp R \ T \ (\x y. Abs x = y)" by (drule Quotient_cr_rel) (auto simp: fun_eq_iff equivp_reflp intro!: eq_reflection) lemma Quotient_crel_typedef: "Quotient (eq_onp P) Abs Rep T \ T \ (\x y. x = Rep y)" unfolding Quotient_def by (auto 0 4 simp: fun_eq_iff eq_onp_def intro: sym intro!: eq_reflection) lemma Quotient_crel_typecopy: "Quotient (=) Abs Rep T \ T \ (\x y. x = Rep y)" by (subst (asm) eq_onp_True[symmetric]) (rule Quotient_crel_typedef) lemma equivp_add_relconj: assumes equiv: "equivp R" "equivp R'" and le: "S OO T OO U \ R OO STU OO R'" shows "R OO S OO T OO U OO R' \ R OO STU OO R'" proof - have trans: "R OO R \ R" "R' OO R' \ R'" using equiv unfolding equivp_reflp_symp_transp transp_relcompp by blast+ have "R OO S OO T OO U OO R' = R OO (S OO T OO U) OO R'" unfolding relcompp_assoc .. also have "\ \ R OO (R OO STU OO R') OO R'" by (intro le relcompp_mono order_refl) also have "\ \ (R OO R) OO STU OO (R' OO R')" unfolding relcompp_assoc .. also have "\ \ R OO STU OO R'" by (intro trans relcompp_mono order_refl) finally show ?thesis . qed lemma Grp_conversep_eq_onp: "((BNF_Def.Grp UNIV f)\\ OO BNF_Def.Grp UNIV f) = eq_onp (\x. x \ range f)" by (auto simp: fun_eq_iff Grp_def eq_onp_def image_iff) lemma Grp_conversep_nonempty: "(BNF_Def.Grp UNIV f)\\ OO BNF_Def.Grp UNIV f \ bot" by (auto simp: fun_eq_iff Grp_def) lemma relcomppI2: "r a b \ s b c \ t c d \ (r OO s OO t) a d" by (auto) lemma rel_conj_eq_onp: "equivp R \ rel_conj R (eq_onp P) \ R" by (auto simp: eq_onp_def transp_def equivp_def) lemma Quotient_Quotient3: "Quotient R Abs Rep T \ Quotient3 R Abs Rep" unfolding Quotient_def Quotient3_def by blast lemma Quotient_reflp_imp_equivp: "Quotient R Abs Rep T \ reflp R \ equivp R" using Quotient_symp Quotient_transp equivpI by blast lemma Quotient_eq_onp_typedef: "Quotient (eq_onp P) Abs Rep cr \ type_definition Rep Abs {x. P x}" unfolding Quotient_def eq_onp_def by unfold_locales auto lemma Quotient_eq_onp_type_copy: "Quotient (=) Abs Rep cr \ type_definition Rep Abs UNIV" unfolding Quotient_def eq_onp_def by unfold_locales auto ML_file "Tools/BNF/bnf_lift.ML" hide_fact sum_insert_Inl_unit lift_sum_unit_vimage_commute insert_Inl_int_map_sum_unit image_map_sum_unit_subset subset_lift_sum_unitD UNIV_sum_unit_conv subset_vimage_image_subset relcompp_mem_Grp_neq_bot comp_projr_Inr in_rel_sum_in_image_projr subset_rel_sumI relcompp_eq_Grp_neq_bot rel_fun_rel_OO1 rel_fun_rel_OO2 rel_sum_eq2_nonempty rel_sum_eq3_nonempty hypsubst equivp_add_relconj Grp_conversep_eq_onp Grp_conversep_nonempty relcomppI2 rel_conj_eq_onp Quotient_reflp_imp_equivp Quotient_Quotient3 end diff --git a/src/HOL/Rings.thy b/src/HOL/Rings.thy --- a/src/HOL/Rings.thy +++ b/src/HOL/Rings.thy @@ -1,2750 +1,2750 @@ (* Title: HOL/Rings.thy Author: Gertrud Bauer Author: Steven Obua Author: Tobias Nipkow Author: Lawrence C Paulson Author: Markus Wenzel Author: Jeremy Avigad *) section \Rings\ theory Rings imports Groups Set Fun begin subsection \Semirings and rings\ class semiring = ab_semigroup_add + semigroup_mult + assumes distrib_right [algebra_simps, algebra_split_simps]: "(a + b) * c = a * c + b * c" assumes distrib_left [algebra_simps, algebra_split_simps]: "a * (b + c) = a * b + a * c" begin text \For the \combine_numerals\ simproc\ lemma combine_common_factor: "a * e + (b * e + c) = (a + b) * e + c" by (simp add: distrib_right ac_simps) end class mult_zero = times + zero + assumes mult_zero_left [simp]: "0 * a = 0" assumes mult_zero_right [simp]: "a * 0 = 0" begin lemma mult_not_zero: "a * b \ 0 \ a \ 0 \ b \ 0" by auto end class semiring_0 = semiring + comm_monoid_add + mult_zero class semiring_0_cancel = semiring + cancel_comm_monoid_add begin subclass semiring_0 proof fix a :: 'a have "0 * a + 0 * a = 0 * a + 0" by (simp add: distrib_right [symmetric]) then show "0 * a = 0" by (simp only: add_left_cancel) have "a * 0 + a * 0 = a * 0 + 0" by (simp add: distrib_left [symmetric]) then show "a * 0 = 0" by (simp only: add_left_cancel) qed end class comm_semiring = ab_semigroup_add + ab_semigroup_mult + assumes distrib: "(a + b) * c = a * c + b * c" begin subclass semiring proof fix a b c :: 'a show "(a + b) * c = a * c + b * c" by (simp add: distrib) have "a * (b + c) = (b + c) * a" by (simp add: ac_simps) also have "\ = b * a + c * a" by (simp only: distrib) also have "\ = a * b + a * c" by (simp add: ac_simps) finally show "a * (b + c) = a * b + a * c" by blast qed end class comm_semiring_0 = comm_semiring + comm_monoid_add + mult_zero begin subclass semiring_0 .. end class comm_semiring_0_cancel = comm_semiring + cancel_comm_monoid_add begin subclass semiring_0_cancel .. subclass comm_semiring_0 .. end class zero_neq_one = zero + one + assumes zero_neq_one [simp]: "0 \ 1" begin lemma one_neq_zero [simp]: "1 \ 0" by (rule not_sym) (rule zero_neq_one) definition of_bool :: "bool \ 'a" where "of_bool p = (if p then 1 else 0)" lemma of_bool_eq [simp, code]: "of_bool False = 0" "of_bool True = 1" by (simp_all add: of_bool_def) lemma of_bool_eq_iff: "of_bool p = of_bool q \ p = q" by (simp add: of_bool_def) lemma split_of_bool [split]: "P (of_bool p) \ (p \ P 1) \ (\ p \ P 0)" by (cases p) simp_all lemma split_of_bool_asm: "P (of_bool p) \ \ (p \ \ P 1 \ \ p \ \ P 0)" by (cases p) simp_all lemma of_bool_eq_0_iff [simp]: \of_bool P = 0 \ \ P\ by simp lemma of_bool_eq_1_iff [simp]: \of_bool P = 1 \ P\ by simp end class semiring_1 = zero_neq_one + semiring_0 + monoid_mult begin lemma of_bool_conj: "of_bool (P \ Q) = of_bool P * of_bool Q" by auto end lemma lambda_zero: "(\h::'a::mult_zero. 0) = (*) 0" by auto lemma lambda_one: "(\x::'a::monoid_mult. x) = (*) 1" by auto subsection \Abstract divisibility\ class dvd = times begin definition dvd :: "'a \ 'a \ bool" (infix "dvd" 50) where "b dvd a \ (\k. a = b * k)" lemma dvdI [intro?]: "a = b * k \ b dvd a" unfolding dvd_def .. lemma dvdE [elim]: "b dvd a \ (\k. a = b * k \ P) \ P" unfolding dvd_def by blast end context comm_monoid_mult begin subclass dvd . lemma dvd_refl [simp]: "a dvd a" proof show "a = a * 1" by simp qed lemma dvd_trans [trans]: assumes "a dvd b" and "b dvd c" shows "a dvd c" proof - from assms obtain v where "b = a * v" by auto moreover from assms obtain w where "c = b * w" by auto ultimately have "c = a * (v * w)" by (simp add: mult.assoc) then show ?thesis .. qed lemma subset_divisors_dvd: "{c. c dvd a} \ {c. c dvd b} \ a dvd b" by (auto simp add: subset_iff intro: dvd_trans) lemma strict_subset_divisors_dvd: "{c. c dvd a} \ {c. c dvd b} \ a dvd b \ \ b dvd a" by (auto simp add: subset_iff intro: dvd_trans) lemma one_dvd [simp]: "1 dvd a" by (auto intro: dvdI) lemma dvd_mult [simp]: "a dvd (b * c)" if "a dvd c" using that by rule (auto intro: mult.left_commute dvdI) lemma dvd_mult2 [simp]: "a dvd (b * c)" if "a dvd b" using that dvd_mult [of a b c] by (simp add: ac_simps) lemma dvd_triv_right [simp]: "a dvd b * a" by (rule dvd_mult) (rule dvd_refl) lemma dvd_triv_left [simp]: "a dvd a * b" by (rule dvd_mult2) (rule dvd_refl) lemma mult_dvd_mono: assumes "a dvd b" and "c dvd d" shows "a * c dvd b * d" proof - from \a dvd b\ obtain b' where "b = a * b'" .. moreover from \c dvd d\ obtain d' where "d = c * d'" .. ultimately have "b * d = (a * c) * (b' * d')" by (simp add: ac_simps) then show ?thesis .. qed lemma dvd_mult_left: "a * b dvd c \ a dvd c" by (simp add: dvd_def mult.assoc) blast lemma dvd_mult_right: "a * b dvd c \ b dvd c" using dvd_mult_left [of b a c] by (simp add: ac_simps) end class comm_semiring_1 = zero_neq_one + comm_semiring_0 + comm_monoid_mult begin subclass semiring_1 .. lemma dvd_0_left_iff [simp]: "0 dvd a \ a = 0" by auto lemma dvd_0_right [iff]: "a dvd 0" proof show "0 = a * 0" by simp qed lemma dvd_0_left: "0 dvd a \ a = 0" by simp lemma dvd_add [simp]: assumes "a dvd b" and "a dvd c" shows "a dvd (b + c)" proof - from \a dvd b\ obtain b' where "b = a * b'" .. moreover from \a dvd c\ obtain c' where "c = a * c'" .. ultimately have "b + c = a * (b' + c')" by (simp add: distrib_left) then show ?thesis .. qed end class semiring_1_cancel = semiring + cancel_comm_monoid_add + zero_neq_one + monoid_mult begin subclass semiring_0_cancel .. subclass semiring_1 .. end class comm_semiring_1_cancel = comm_semiring + cancel_comm_monoid_add + zero_neq_one + comm_monoid_mult + assumes right_diff_distrib' [algebra_simps, algebra_split_simps]: "a * (b - c) = a * b - a * c" begin subclass semiring_1_cancel .. subclass comm_semiring_0_cancel .. subclass comm_semiring_1 .. lemma left_diff_distrib' [algebra_simps, algebra_split_simps]: "(b - c) * a = b * a - c * a" by (simp add: algebra_simps) lemma dvd_add_times_triv_left_iff [simp]: "a dvd c * a + b \ a dvd b" proof - have "a dvd a * c + b \ a dvd b" (is "?P \ ?Q") proof assume ?Q then show ?P by simp next assume ?P then obtain d where "a * c + b = a * d" .. then have "a * c + b - a * c = a * d - a * c" by simp then have "b = a * d - a * c" by simp then have "b = a * (d - c)" by (simp add: algebra_simps) then show ?Q .. qed then show "a dvd c * a + b \ a dvd b" by (simp add: ac_simps) qed lemma dvd_add_times_triv_right_iff [simp]: "a dvd b + c * a \ a dvd b" using dvd_add_times_triv_left_iff [of a c b] by (simp add: ac_simps) lemma dvd_add_triv_left_iff [simp]: "a dvd a + b \ a dvd b" using dvd_add_times_triv_left_iff [of a 1 b] by simp lemma dvd_add_triv_right_iff [simp]: "a dvd b + a \ a dvd b" using dvd_add_times_triv_right_iff [of a b 1] by simp lemma dvd_add_right_iff: assumes "a dvd b" shows "a dvd b + c \ a dvd c" (is "?P \ ?Q") proof assume ?P then obtain d where "b + c = a * d" .. moreover from \a dvd b\ obtain e where "b = a * e" .. ultimately have "a * e + c = a * d" by simp then have "a * e + c - a * e = a * d - a * e" by simp then have "c = a * d - a * e" by simp then have "c = a * (d - e)" by (simp add: algebra_simps) then show ?Q .. next assume ?Q with assms show ?P by simp qed lemma dvd_add_left_iff: "a dvd c \ a dvd b + c \ a dvd b" using dvd_add_right_iff [of a c b] by (simp add: ac_simps) end class ring = semiring + ab_group_add begin subclass semiring_0_cancel .. text \Distribution rules\ lemma minus_mult_left: "- (a * b) = - a * b" by (rule minus_unique) (simp add: distrib_right [symmetric]) lemma minus_mult_right: "- (a * b) = a * - b" by (rule minus_unique) (simp add: distrib_left [symmetric]) text \Extract signs from products\ lemmas mult_minus_left [simp] = minus_mult_left [symmetric] lemmas mult_minus_right [simp] = minus_mult_right [symmetric] lemma minus_mult_minus [simp]: "- a * - b = a * b" by simp lemma minus_mult_commute: "- a * b = a * - b" by simp lemma right_diff_distrib [algebra_simps, algebra_split_simps]: "a * (b - c) = a * b - a * c" using distrib_left [of a b "-c "] by simp lemma left_diff_distrib [algebra_simps, algebra_split_simps]: "(a - b) * c = a * c - b * c" using distrib_right [of a "- b" c] by simp lemmas ring_distribs = distrib_left distrib_right left_diff_distrib right_diff_distrib lemma eq_add_iff1: "a * e + c = b * e + d \ (a - b) * e + c = d" by (simp add: algebra_simps) lemma eq_add_iff2: "a * e + c = b * e + d \ c = (b - a) * e + d" by (simp add: algebra_simps) end lemmas ring_distribs = distrib_left distrib_right left_diff_distrib right_diff_distrib class comm_ring = comm_semiring + ab_group_add begin subclass ring .. subclass comm_semiring_0_cancel .. lemma square_diff_square_factored: "x * x - y * y = (x + y) * (x - y)" by (simp add: algebra_simps) end class ring_1 = ring + zero_neq_one + monoid_mult begin subclass semiring_1_cancel .. lemma of_bool_not_iff [simp]: \of_bool (\ P) = 1 - of_bool P\ by simp lemma square_diff_one_factored: "x * x - 1 = (x + 1) * (x - 1)" by (simp add: algebra_simps) end class comm_ring_1 = comm_ring + zero_neq_one + comm_monoid_mult begin subclass ring_1 .. subclass comm_semiring_1_cancel by standard (simp add: algebra_simps) lemma dvd_minus_iff [simp]: "x dvd - y \ x dvd y" proof assume "x dvd - y" then have "x dvd - 1 * - y" by (rule dvd_mult) then show "x dvd y" by simp next assume "x dvd y" then have "x dvd - 1 * y" by (rule dvd_mult) then show "x dvd - y" by simp qed lemma minus_dvd_iff [simp]: "- x dvd y \ x dvd y" proof assume "- x dvd y" then obtain k where "y = - x * k" .. then have "y = x * - k" by simp then show "x dvd y" .. next assume "x dvd y" then obtain k where "y = x * k" .. then have "y = - x * - k" by simp then show "- x dvd y" .. qed lemma dvd_diff [simp]: "x dvd y \ x dvd z \ x dvd (y - z)" using dvd_add [of x y "- z"] by simp end subsection \Towards integral domains\ class semiring_no_zero_divisors = semiring_0 + assumes no_zero_divisors: "a \ 0 \ b \ 0 \ a * b \ 0" begin lemma divisors_zero: assumes "a * b = 0" shows "a = 0 \ b = 0" proof (rule classical) assume "\ ?thesis" then have "a \ 0" and "b \ 0" by auto with no_zero_divisors have "a * b \ 0" by blast with assms show ?thesis by simp qed lemma mult_eq_0_iff [simp]: "a * b = 0 \ a = 0 \ b = 0" proof (cases "a = 0 \ b = 0") case False then have "a \ 0" and "b \ 0" by auto then show ?thesis using no_zero_divisors by simp next case True then show ?thesis by auto qed end class semiring_1_no_zero_divisors = semiring_1 + semiring_no_zero_divisors class semiring_no_zero_divisors_cancel = semiring_no_zero_divisors + assumes mult_cancel_right [simp]: "a * c = b * c \ c = 0 \ a = b" and mult_cancel_left [simp]: "c * a = c * b \ c = 0 \ a = b" begin lemma mult_left_cancel: "c \ 0 \ c * a = c * b \ a = b" by simp lemma mult_right_cancel: "c \ 0 \ a * c = b * c \ a = b" by simp end class ring_no_zero_divisors = ring + semiring_no_zero_divisors begin subclass semiring_no_zero_divisors_cancel proof fix a b c have "a * c = b * c \ (a - b) * c = 0" by (simp add: algebra_simps) also have "\ \ c = 0 \ a = b" by auto finally show "a * c = b * c \ c = 0 \ a = b" . have "c * a = c * b \ c * (a - b) = 0" by (simp add: algebra_simps) also have "\ \ c = 0 \ a = b" by auto finally show "c * a = c * b \ c = 0 \ a = b" . qed end class ring_1_no_zero_divisors = ring_1 + ring_no_zero_divisors begin subclass semiring_1_no_zero_divisors .. lemma square_eq_1_iff: "x * x = 1 \ x = 1 \ x = - 1" proof - have "(x - 1) * (x + 1) = x * x - 1" by (simp add: algebra_simps) then have "x * x = 1 \ (x - 1) * (x + 1) = 0" by simp then show ?thesis by (simp add: eq_neg_iff_add_eq_0) qed lemma mult_cancel_right1 [simp]: "c = b * c \ c = 0 \ b = 1" using mult_cancel_right [of 1 c b] by auto lemma mult_cancel_right2 [simp]: "a * c = c \ c = 0 \ a = 1" using mult_cancel_right [of a c 1] by simp lemma mult_cancel_left1 [simp]: "c = c * b \ c = 0 \ b = 1" using mult_cancel_left [of c 1 b] by force lemma mult_cancel_left2 [simp]: "c * a = c \ c = 0 \ a = 1" using mult_cancel_left [of c a 1] by simp end class semidom = comm_semiring_1_cancel + semiring_no_zero_divisors begin subclass semiring_1_no_zero_divisors .. end class idom = comm_ring_1 + semiring_no_zero_divisors begin subclass semidom .. subclass ring_1_no_zero_divisors .. lemma dvd_mult_cancel_right [simp]: "a * c dvd b * c \ c = 0 \ a dvd b" proof - have "a * c dvd b * c \ (\k. b * c = (a * k) * c)" by (auto simp add: ac_simps) also have "(\k. b * c = (a * k) * c) \ c = 0 \ a dvd b" by auto finally show ?thesis . qed lemma dvd_mult_cancel_left [simp]: "c * a dvd c * b \ c = 0 \ a dvd b" using dvd_mult_cancel_right [of a c b] by (simp add: ac_simps) lemma square_eq_iff: "a * a = b * b \ a = b \ a = - b" proof assume "a * a = b * b" then have "(a - b) * (a + b) = 0" by (simp add: algebra_simps) then show "a = b \ a = - b" by (simp add: eq_neg_iff_add_eq_0) next assume "a = b \ a = - b" then show "a * a = b * b" by auto qed end class idom_abs_sgn = idom + abs + sgn + assumes sgn_mult_abs: "sgn a * \a\ = a" and sgn_sgn [simp]: "sgn (sgn a) = sgn a" and abs_abs [simp]: "\\a\\ = \a\" and abs_0 [simp]: "\0\ = 0" and sgn_0 [simp]: "sgn 0 = 0" and sgn_1 [simp]: "sgn 1 = 1" and sgn_minus_1: "sgn (- 1) = - 1" and sgn_mult: "sgn (a * b) = sgn a * sgn b" begin lemma sgn_eq_0_iff: "sgn a = 0 \ a = 0" proof - { assume "sgn a = 0" then have "sgn a * \a\ = 0" by simp then have "a = 0" by (simp add: sgn_mult_abs) } then show ?thesis by auto qed lemma abs_eq_0_iff: "\a\ = 0 \ a = 0" proof - { assume "\a\ = 0" then have "sgn a * \a\ = 0" by simp then have "a = 0" by (simp add: sgn_mult_abs) } then show ?thesis by auto qed lemma abs_mult_sgn: "\a\ * sgn a = a" using sgn_mult_abs [of a] by (simp add: ac_simps) lemma abs_1 [simp]: "\1\ = 1" using sgn_mult_abs [of 1] by simp lemma sgn_abs [simp]: "\sgn a\ = of_bool (a \ 0)" using sgn_mult_abs [of "sgn a"] mult_cancel_left [of "sgn a" "\sgn a\" 1] by (auto simp add: sgn_eq_0_iff) lemma abs_sgn [simp]: "sgn \a\ = of_bool (a \ 0)" using sgn_mult_abs [of "\a\"] mult_cancel_right [of "sgn \a\" "\a\" 1] by (auto simp add: abs_eq_0_iff) lemma abs_mult: "\a * b\ = \a\ * \b\" proof (cases "a = 0 \ b = 0") case True then show ?thesis by auto next case False then have *: "sgn (a * b) \ 0" by (simp add: sgn_eq_0_iff) from abs_mult_sgn [of "a * b"] abs_mult_sgn [of a] abs_mult_sgn [of b] have "\a * b\ * sgn (a * b) = \a\ * sgn a * \b\ * sgn b" by (simp add: ac_simps) then have "\a * b\ * sgn (a * b) = \a\ * \b\ * sgn (a * b)" by (simp add: sgn_mult ac_simps) with * show ?thesis by simp qed lemma sgn_minus [simp]: "sgn (- a) = - sgn a" proof - from sgn_minus_1 have "sgn (- 1 * a) = - 1 * sgn a" by (simp only: sgn_mult) then show ?thesis by simp qed lemma abs_minus [simp]: "\- a\ = \a\" proof - have [simp]: "\- 1\ = 1" using sgn_mult_abs [of "- 1"] by simp then have "\- 1 * a\ = 1 * \a\" by (simp only: abs_mult) then show ?thesis by simp qed end subsection \(Partial) Division\ class divide = fixes divide :: "'a \ 'a \ 'a" (infixl "div" 70) setup \Sign.add_const_constraint (\<^const_name>\divide\, SOME \<^typ>\'a \ 'a \ 'a\)\ context semiring begin lemma [field_simps, field_split_simps]: shows distrib_left_NO_MATCH: "NO_MATCH (x div y) a \ a * (b + c) = a * b + a * c" and distrib_right_NO_MATCH: "NO_MATCH (x div y) c \ (a + b) * c = a * c + b * c" by (rule distrib_left distrib_right)+ end context ring begin lemma [field_simps, field_split_simps]: shows left_diff_distrib_NO_MATCH: "NO_MATCH (x div y) c \ (a - b) * c = a * c - b * c" and right_diff_distrib_NO_MATCH: "NO_MATCH (x div y) a \ a * (b - c) = a * b - a * c" by (rule left_diff_distrib right_diff_distrib)+ end setup \Sign.add_const_constraint (\<^const_name>\divide\, SOME \<^typ>\'a::divide \ 'a \ 'a\)\ text \Algebraic classes with division\ class semidom_divide = semidom + divide + assumes nonzero_mult_div_cancel_right [simp]: "b \ 0 \ (a * b) div b = a" assumes div_by_0 [simp]: "a div 0 = 0" begin lemma nonzero_mult_div_cancel_left [simp]: "a \ 0 \ (a * b) div a = b" using nonzero_mult_div_cancel_right [of a b] by (simp add: ac_simps) subclass semiring_no_zero_divisors_cancel proof show *: "a * c = b * c \ c = 0 \ a = b" for a b c proof (cases "c = 0") case True then show ?thesis by simp next case False have "a = b" if "a * c = b * c" proof - from that have "a * c div c = b * c div c" by simp with False show ?thesis by simp qed then show ?thesis by auto qed show "c * a = c * b \ c = 0 \ a = b" for a b c using * [of a c b] by (simp add: ac_simps) qed lemma div_self [simp]: "a \ 0 \ a div a = 1" using nonzero_mult_div_cancel_left [of a 1] by simp lemma div_0 [simp]: "0 div a = 0" proof (cases "a = 0") case True then show ?thesis by simp next case False then have "a * 0 div a = 0" by (rule nonzero_mult_div_cancel_left) then show ?thesis by simp qed lemma div_by_1 [simp]: "a div 1 = a" using nonzero_mult_div_cancel_left [of 1 a] by simp lemma dvd_div_eq_0_iff: assumes "b dvd a" shows "a div b = 0 \ a = 0" using assms by (elim dvdE, cases "b = 0") simp_all lemma dvd_div_eq_cancel: "a div c = b div c \ c dvd a \ c dvd b \ a = b" by (elim dvdE, cases "c = 0") simp_all lemma dvd_div_eq_iff: "c dvd a \ c dvd b \ a div c = b div c \ a = b" by (elim dvdE, cases "c = 0") simp_all lemma inj_on_mult: "inj_on ((*) a) A" if "a \ 0" proof (rule inj_onI) fix b c assume "a * b = a * c" then have "a * b div a = a * c div a" by (simp only:) with that show "b = c" by simp qed end class idom_divide = idom + semidom_divide begin lemma dvd_neg_div: assumes "b dvd a" shows "- a div b = - (a div b)" proof (cases "b = 0") case True then show ?thesis by simp next case False from assms obtain c where "a = b * c" .. then have "- a div b = (b * - c) div b" by simp from False also have "\ = - c" by (rule nonzero_mult_div_cancel_left) with False \a = b * c\ show ?thesis by simp qed lemma dvd_div_neg: assumes "b dvd a" shows "a div - b = - (a div b)" proof (cases "b = 0") case True then show ?thesis by simp next case False then have "- b \ 0" by simp from assms obtain c where "a = b * c" .. then have "a div - b = (- b * - c) div - b" by simp from \- b \ 0\ also have "\ = - c" by (rule nonzero_mult_div_cancel_left) with False \a = b * c\ show ?thesis by simp qed end class algebraic_semidom = semidom_divide begin text \ Class \<^class>\algebraic_semidom\ enriches a integral domain by notions from algebra, like units in a ring. It is a separate class to avoid spoiling fields with notions which are degenerated there. \ lemma dvd_times_left_cancel_iff [simp]: assumes "a \ 0" shows "a * b dvd a * c \ b dvd c" (is "?lhs \ ?rhs") proof assume ?lhs then obtain d where "a * c = a * b * d" .. with assms have "c = b * d" by (simp add: ac_simps) then show ?rhs .. next assume ?rhs then obtain d where "c = b * d" .. then have "a * c = a * b * d" by (simp add: ac_simps) then show ?lhs .. qed lemma dvd_times_right_cancel_iff [simp]: assumes "a \ 0" shows "b * a dvd c * a \ b dvd c" using dvd_times_left_cancel_iff [of a b c] assms by (simp add: ac_simps) lemma div_dvd_iff_mult: assumes "b \ 0" and "b dvd a" shows "a div b dvd c \ a dvd c * b" proof - from \b dvd a\ obtain d where "a = b * d" .. with \b \ 0\ show ?thesis by (simp add: ac_simps) qed lemma dvd_div_iff_mult: assumes "c \ 0" and "c dvd b" shows "a dvd b div c \ a * c dvd b" proof - from \c dvd b\ obtain d where "b = c * d" .. with \c \ 0\ show ?thesis by (simp add: mult.commute [of a]) qed lemma div_dvd_div [simp]: assumes "a dvd b" and "a dvd c" shows "b div a dvd c div a \ b dvd c" proof (cases "a = 0") case True with assms show ?thesis by simp next case False moreover from assms obtain k l where "b = a * k" and "c = a * l" by blast ultimately show ?thesis by simp qed lemma div_add [simp]: assumes "c dvd a" and "c dvd b" shows "(a + b) div c = a div c + b div c" proof (cases "c = 0") case True then show ?thesis by simp next case False moreover from assms obtain k l where "a = c * k" and "b = c * l" by blast moreover have "c * k + c * l = c * (k + l)" by (simp add: algebra_simps) ultimately show ?thesis by simp qed lemma div_mult_div_if_dvd: assumes "b dvd a" and "d dvd c" shows "(a div b) * (c div d) = (a * c) div (b * d)" proof (cases "b = 0 \ c = 0") case True with assms show ?thesis by auto next case False moreover from assms obtain k l where "a = b * k" and "c = d * l" by blast moreover have "b * k * (d * l) div (b * d) = (b * d) * (k * l) div (b * d)" by (simp add: ac_simps) ultimately show ?thesis by simp qed lemma dvd_div_eq_mult: assumes "a \ 0" and "a dvd b" shows "b div a = c \ b = c * a" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by (simp add: assms) next assume ?lhs then have "b div a * a = c * a" by simp moreover from assms have "b div a * a = b" by (auto simp add: ac_simps) ultimately show ?rhs by simp qed lemma dvd_div_mult_self [simp]: "a dvd b \ b div a * a = b" by (cases "a = 0") (auto simp add: ac_simps) lemma dvd_mult_div_cancel [simp]: "a dvd b \ a * (b div a) = b" using dvd_div_mult_self [of a b] by (simp add: ac_simps) lemma div_mult_swap: assumes "c dvd b" shows "a * (b div c) = (a * b) div c" proof (cases "c = 0") case True then show ?thesis by simp next case False from assms obtain d where "b = c * d" .. moreover from False have "a * divide (d * c) c = ((a * d) * c) div c" by simp ultimately show ?thesis by (simp add: ac_simps) qed lemma dvd_div_mult: "c dvd b \ b div c * a = (b * a) div c" using div_mult_swap [of c b a] by (simp add: ac_simps) lemma dvd_div_mult2_eq: assumes "b * c dvd a" shows "a div (b * c) = a div b div c" proof - from assms obtain k where "a = b * c * k" .. then show ?thesis by (cases "b = 0 \ c = 0") (auto, simp add: ac_simps) qed lemma dvd_div_div_eq_mult: assumes "a \ 0" "c \ 0" and "a dvd b" "c dvd d" shows "b div a = d div c \ b * c = a * d" (is "?lhs \ ?rhs") proof - from assms have "a * c \ 0" by simp then have "?lhs \ b div a * (a * c) = d div c * (a * c)" by simp also have "\ \ (a * (b div a)) * c = (c * (d div c)) * a" by (simp add: ac_simps) also have "\ \ (a * b div a) * c = (c * d div c) * a" using assms by (simp add: div_mult_swap) also have "\ \ ?rhs" using assms by (simp add: ac_simps) finally show ?thesis . qed lemma dvd_mult_imp_div: assumes "a * c dvd b" shows "a dvd b div c" proof (cases "c = 0") case True then show ?thesis by simp next case False from \a * c dvd b\ obtain d where "b = a * c * d" .. with False show ?thesis by (simp add: mult.commute [of a] mult.assoc) qed lemma div_div_eq_right: assumes "c dvd b" "b dvd a" shows "a div (b div c) = a div b * c" proof (cases "c = 0 \ b = 0") case True then show ?thesis by auto next case False from assms obtain r s where "b = c * r" and "a = c * r * s" by blast moreover with False have "r \ 0" by auto ultimately show ?thesis using False by simp (simp add: mult.commute [of _ r] mult.assoc mult.commute [of c]) qed lemma div_div_div_same: assumes "d dvd b" "b dvd a" shows "(a div d) div (b div d) = a div b" proof (cases "b = 0 \ d = 0") case True with assms show ?thesis by auto next case False from assms obtain r s where "a = d * r * s" and "b = d * r" by blast with False show ?thesis by simp (simp add: ac_simps) qed text \Units: invertible elements in a ring\ abbreviation is_unit :: "'a \ bool" where "is_unit a \ a dvd 1" lemma not_is_unit_0 [simp]: "\ is_unit 0" by simp lemma unit_imp_dvd [dest]: "is_unit b \ b dvd a" by (rule dvd_trans [of _ 1]) simp_all lemma unit_dvdE: assumes "is_unit a" obtains c where "a \ 0" and "b = a * c" proof - from assms have "a dvd b" by auto then obtain c where "b = a * c" .. moreover from assms have "a \ 0" by auto ultimately show thesis using that by blast qed lemma dvd_unit_imp_unit: "a dvd b \ is_unit b \ is_unit a" by (rule dvd_trans) lemma unit_div_1_unit [simp, intro]: assumes "is_unit a" shows "is_unit (1 div a)" proof - from assms have "1 = 1 div a * a" by simp then show "is_unit (1 div a)" by (rule dvdI) qed lemma is_unitE [elim?]: assumes "is_unit a" obtains b where "a \ 0" and "b \ 0" and "is_unit b" and "1 div a = b" and "1 div b = a" and "a * b = 1" and "c div a = c * b" proof (rule that) define b where "b = 1 div a" then show "1 div a = b" by simp from assms b_def show "is_unit b" by simp with assms show "a \ 0" and "b \ 0" by auto from assms b_def show "a * b = 1" by simp then have "1 = a * b" .. with b_def \b \ 0\ show "1 div b = a" by simp from assms have "a dvd c" .. then obtain d where "c = a * d" .. with \a \ 0\ \a * b = 1\ show "c div a = c * b" by (simp add: mult.assoc mult.left_commute [of a]) qed lemma unit_prod [intro]: "is_unit a \ is_unit b \ is_unit (a * b)" by (subst mult_1_left [of 1, symmetric]) (rule mult_dvd_mono) lemma is_unit_mult_iff: "is_unit (a * b) \ is_unit a \ is_unit b" by (auto dest: dvd_mult_left dvd_mult_right) lemma unit_div [intro]: "is_unit a \ is_unit b \ is_unit (a div b)" by (erule is_unitE [of b a]) (simp add: ac_simps unit_prod) lemma mult_unit_dvd_iff: assumes "is_unit b" shows "a * b dvd c \ a dvd c" proof assume "a * b dvd c" with assms show "a dvd c" by (simp add: dvd_mult_left) next assume "a dvd c" then obtain k where "c = a * k" .. with assms have "c = (a * b) * (1 div b * k)" by (simp add: mult_ac) then show "a * b dvd c" by (rule dvdI) qed lemma mult_unit_dvd_iff': "is_unit a \ (a * b) dvd c \ b dvd c" using mult_unit_dvd_iff [of a b c] by (simp add: ac_simps) lemma dvd_mult_unit_iff: assumes "is_unit b" shows "a dvd c * b \ a dvd c" proof assume "a dvd c * b" with assms have "c * b dvd c * (b * (1 div b))" by (subst mult_assoc [symmetric]) simp also from assms have "b * (1 div b) = 1" by (rule is_unitE) simp finally have "c * b dvd c" by simp with \a dvd c * b\ show "a dvd c" by (rule dvd_trans) next assume "a dvd c" then show "a dvd c * b" by simp qed lemma dvd_mult_unit_iff': "is_unit b \ a dvd b * c \ a dvd c" using dvd_mult_unit_iff [of b a c] by (simp add: ac_simps) lemma div_unit_dvd_iff: "is_unit b \ a div b dvd c \ a dvd c" by (erule is_unitE [of _ a]) (auto simp add: mult_unit_dvd_iff) lemma dvd_div_unit_iff: "is_unit b \ a dvd c div b \ a dvd c" by (erule is_unitE [of _ c]) (simp add: dvd_mult_unit_iff) lemmas unit_dvd_iff = mult_unit_dvd_iff mult_unit_dvd_iff' dvd_mult_unit_iff dvd_mult_unit_iff' div_unit_dvd_iff dvd_div_unit_iff (* FIXME consider named_theorems *) lemma unit_mult_div_div [simp]: "is_unit a \ b * (1 div a) = b div a" by (erule is_unitE [of _ b]) simp lemma unit_div_mult_self [simp]: "is_unit a \ b div a * a = b" by (rule dvd_div_mult_self) auto lemma unit_div_1_div_1 [simp]: "is_unit a \ 1 div (1 div a) = a" by (erule is_unitE) simp lemma unit_div_mult_swap: "is_unit c \ a * (b div c) = (a * b) div c" by (erule unit_dvdE [of _ b]) (simp add: mult.left_commute [of _ c]) lemma unit_div_commute: "is_unit b \ (a div b) * c = (a * c) div b" using unit_div_mult_swap [of b c a] by (simp add: ac_simps) lemma unit_eq_div1: "is_unit b \ a div b = c \ a = c * b" by (auto elim: is_unitE) lemma unit_eq_div2: "is_unit b \ a = c div b \ a * b = c" using unit_eq_div1 [of b c a] by auto lemma unit_mult_left_cancel: "is_unit a \ a * b = a * c \ b = c" using mult_cancel_left [of a b c] by auto lemma unit_mult_right_cancel: "is_unit a \ b * a = c * a \ b = c" using unit_mult_left_cancel [of a b c] by (auto simp add: ac_simps) lemma unit_div_cancel: assumes "is_unit a" shows "b div a = c div a \ b = c" proof - from assms have "is_unit (1 div a)" by simp then have "b * (1 div a) = c * (1 div a) \ b = c" by (rule unit_mult_right_cancel) with assms show ?thesis by simp qed lemma is_unit_div_mult2_eq: assumes "is_unit b" and "is_unit c" shows "a div (b * c) = a div b div c" proof - from assms have "is_unit (b * c)" by (simp add: unit_prod) then have "b * c dvd a" by (rule unit_imp_dvd) then show ?thesis by (rule dvd_div_mult2_eq) qed lemma is_unit_div_mult_cancel_left: assumes "a \ 0" and "is_unit b" shows "a div (a * b) = 1 div b" proof - from assms have "a div (a * b) = a div a div b" by (simp add: mult_unit_dvd_iff dvd_div_mult2_eq) with assms show ?thesis by simp qed lemma is_unit_div_mult_cancel_right: assumes "a \ 0" and "is_unit b" shows "a div (b * a) = 1 div b" using assms is_unit_div_mult_cancel_left [of a b] by (simp add: ac_simps) lemma unit_div_eq_0_iff: assumes "is_unit b" shows "a div b = 0 \ a = 0" by (rule dvd_div_eq_0_iff) (insert assms, auto) lemma div_mult_unit2: "is_unit c \ b dvd a \ a div (b * c) = a div b div c" by (rule dvd_div_mult2_eq) (simp_all add: mult_unit_dvd_iff) text \Coprimality\ definition coprime :: "'a \ 'a \ bool" where "coprime a b \ (\c. c dvd a \ c dvd b \ is_unit c)" lemma coprimeI: assumes "\c. c dvd a \ c dvd b \ is_unit c" shows "coprime a b" using assms by (auto simp: coprime_def) lemma not_coprimeI: assumes "c dvd a" and "c dvd b" and "\ is_unit c" shows "\ coprime a b" using assms by (auto simp: coprime_def) lemma coprime_common_divisor: "is_unit c" if "coprime a b" and "c dvd a" and "c dvd b" using that by (auto simp: coprime_def) lemma not_coprimeE: assumes "\ coprime a b" obtains c where "c dvd a" and "c dvd b" and "\ is_unit c" using assms by (auto simp: coprime_def) lemma coprime_imp_coprime: "coprime a b" if "coprime c d" and "\e. \ is_unit e \ e dvd a \ e dvd b \ e dvd c" and "\e. \ is_unit e \ e dvd a \ e dvd b \ e dvd d" proof (rule coprimeI) fix e assume "e dvd a" and "e dvd b" with that have "e dvd c" and "e dvd d" by (auto intro: dvd_trans) with \coprime c d\ show "is_unit e" by (rule coprime_common_divisor) qed lemma coprime_divisors: "coprime a b" if "a dvd c" "b dvd d" and "coprime c d" using \coprime c d\ proof (rule coprime_imp_coprime) fix e assume "e dvd a" then show "e dvd c" using \a dvd c\ by (rule dvd_trans) assume "e dvd b" then show "e dvd d" using \b dvd d\ by (rule dvd_trans) qed lemma coprime_self [simp]: "coprime a a \ is_unit a" (is "?P \ ?Q") proof assume ?P then show ?Q by (rule coprime_common_divisor) simp_all next assume ?Q show ?P by (rule coprimeI) (erule dvd_unit_imp_unit, rule \?Q\) qed lemma coprime_commute [ac_simps]: "coprime b a \ coprime a b" unfolding coprime_def by auto lemma is_unit_left_imp_coprime: "coprime a b" if "is_unit a" proof (rule coprimeI) fix c assume "c dvd a" with that show "is_unit c" by (auto intro: dvd_unit_imp_unit) qed lemma is_unit_right_imp_coprime: "coprime a b" if "is_unit b" using that is_unit_left_imp_coprime [of b a] by (simp add: ac_simps) lemma coprime_1_left [simp]: "coprime 1 a" by (rule coprimeI) lemma coprime_1_right [simp]: "coprime a 1" by (rule coprimeI) lemma coprime_0_left_iff [simp]: "coprime 0 a \ is_unit a" by (auto intro: coprimeI dvd_unit_imp_unit coprime_common_divisor [of 0 a a]) lemma coprime_0_right_iff [simp]: "coprime a 0 \ is_unit a" using coprime_0_left_iff [of a] by (simp add: ac_simps) lemma coprime_mult_self_left_iff [simp]: "coprime (c * a) (c * b) \ is_unit c \ coprime a b" by (auto intro: coprime_common_divisor) (rule coprimeI, auto intro: coprime_common_divisor simp add: dvd_mult_unit_iff')+ lemma coprime_mult_self_right_iff [simp]: "coprime (a * c) (b * c) \ is_unit c \ coprime a b" using coprime_mult_self_left_iff [of c a b] by (simp add: ac_simps) lemma coprime_absorb_left: assumes "x dvd y" shows "coprime x y \ is_unit x" using assms coprime_common_divisor is_unit_left_imp_coprime by auto lemma coprime_absorb_right: assumes "y dvd x" shows "coprime x y \ is_unit y" using assms coprime_common_divisor is_unit_right_imp_coprime by auto end class unit_factor = fixes unit_factor :: "'a \ 'a" class semidom_divide_unit_factor = semidom_divide + unit_factor + assumes unit_factor_0 [simp]: "unit_factor 0 = 0" and is_unit_unit_factor: "a dvd 1 \ unit_factor a = a" and unit_factor_is_unit: "a \ 0 \ unit_factor a dvd 1" and unit_factor_mult_unit_left: "a dvd 1 \ unit_factor (a * b) = a * unit_factor b" \ \This fine-grained hierarchy will later on allow lean normalization of polynomials\ begin lemma unit_factor_mult_unit_right: "a dvd 1 \ unit_factor (b * a) = unit_factor b * a" using unit_factor_mult_unit_left[of a b] by (simp add: mult_ac) lemmas [simp] = unit_factor_mult_unit_left unit_factor_mult_unit_right end class normalization_semidom = algebraic_semidom + semidom_divide_unit_factor + fixes normalize :: "'a \ 'a" assumes unit_factor_mult_normalize [simp]: "unit_factor a * normalize a = a" and normalize_0 [simp]: "normalize 0 = 0" begin text \ Class \<^class>\normalization_semidom\ cultivates the idea that each integral domain can be split into equivalence classes whose representants are associated, i.e. divide each other. \<^const>\normalize\ specifies a canonical representant for each equivalence class. The rationale behind this is that it is easier to reason about equality than equivalences, hence we prefer to think about equality of normalized values rather than associated elements. \ declare unit_factor_is_unit [iff] lemma unit_factor_dvd [simp]: "a \ 0 \ unit_factor a dvd b" by (rule unit_imp_dvd) simp lemma unit_factor_self [simp]: "unit_factor a dvd a" by (cases "a = 0") simp_all lemma normalize_mult_unit_factor [simp]: "normalize a * unit_factor a = a" using unit_factor_mult_normalize [of a] by (simp add: ac_simps) lemma normalize_eq_0_iff [simp]: "normalize a = 0 \ a = 0" (is "?lhs \ ?rhs") proof assume ?lhs moreover have "unit_factor a * normalize a = a" by simp ultimately show ?rhs by simp next assume ?rhs then show ?lhs by simp qed lemma unit_factor_eq_0_iff [simp]: "unit_factor a = 0 \ a = 0" (is "?lhs \ ?rhs") proof assume ?lhs moreover have "unit_factor a * normalize a = a" by simp ultimately show ?rhs by simp next assume ?rhs then show ?lhs by simp qed lemma div_unit_factor [simp]: "a div unit_factor a = normalize a" proof (cases "a = 0") case True then show ?thesis by simp next case False then have "unit_factor a \ 0" by simp with nonzero_mult_div_cancel_left have "unit_factor a * normalize a div unit_factor a = normalize a" by blast then show ?thesis by simp qed lemma normalize_div [simp]: "normalize a div a = 1 div unit_factor a" proof (cases "a = 0") case True then show ?thesis by simp next case False have "normalize a div a = normalize a div (unit_factor a * normalize a)" by simp also have "\ = 1 div unit_factor a" using False by (subst is_unit_div_mult_cancel_right) simp_all finally show ?thesis . qed lemma is_unit_normalize: assumes "is_unit a" shows "normalize a = 1" proof - from assms have "unit_factor a = a" by (rule is_unit_unit_factor) moreover from assms have "a \ 0" by auto moreover have "normalize a = a div unit_factor a" by simp ultimately show ?thesis by simp qed lemma unit_factor_1 [simp]: "unit_factor 1 = 1" by (rule is_unit_unit_factor) simp lemma normalize_1 [simp]: "normalize 1 = 1" by (rule is_unit_normalize) simp lemma normalize_1_iff: "normalize a = 1 \ is_unit a" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by (rule is_unit_normalize) next assume ?lhs then have "unit_factor a * normalize a = unit_factor a * 1" by simp then have "unit_factor a = a" by simp moreover from \?lhs\ have "a \ 0" by auto then have "is_unit (unit_factor a)" by simp ultimately show ?rhs by simp qed lemma div_normalize [simp]: "a div normalize a = unit_factor a" proof (cases "a = 0") case True then show ?thesis by simp next case False then have "normalize a \ 0" by simp with nonzero_mult_div_cancel_right have "unit_factor a * normalize a div normalize a = unit_factor a" by blast then show ?thesis by simp qed lemma mult_one_div_unit_factor [simp]: "a * (1 div unit_factor b) = a div unit_factor b" by (cases "b = 0") simp_all lemma inv_unit_factor_eq_0_iff [simp]: "1 div unit_factor a = 0 \ a = 0" (is "?lhs \ ?rhs") proof assume ?lhs then have "a * (1 div unit_factor a) = a * 0" by simp then show ?rhs by simp next assume ?rhs then show ?lhs by simp qed lemma unit_factor_idem [simp]: "unit_factor (unit_factor a) = unit_factor a" by (cases "a = 0") (auto intro: is_unit_unit_factor) lemma normalize_unit_factor [simp]: "a \ 0 \ normalize (unit_factor a) = 1" by (rule is_unit_normalize) simp lemma normalize_mult_unit_left [simp]: assumes "a dvd 1" shows "normalize (a * b) = normalize b" proof (cases "b = 0") case False have "a * unit_factor b * normalize (a * b) = unit_factor (a * b) * normalize (a * b)" using assms by (subst unit_factor_mult_unit_left) auto also have "\ = a * b" by simp also have "b = unit_factor b * normalize b" by simp hence "a * b = a * unit_factor b * normalize b" by (simp only: mult_ac) finally show ?thesis using assms False by auto qed auto lemma normalize_mult_unit_right [simp]: assumes "b dvd 1" shows "normalize (a * b) = normalize a" using assms by (subst mult.commute) auto lemma normalize_idem [simp]: "normalize (normalize a) = normalize a" proof (cases "a = 0") case False have "normalize a = normalize (unit_factor a * normalize a)" by simp also from False have "\ = normalize (normalize a)" by (subst normalize_mult_unit_left) auto finally show ?thesis .. qed auto lemma unit_factor_normalize [simp]: assumes "a \ 0" shows "unit_factor (normalize a) = 1" proof - from assms have *: "normalize a \ 0" by simp have "unit_factor (normalize a) * normalize (normalize a) = normalize a" by (simp only: unit_factor_mult_normalize) then have "unit_factor (normalize a) * normalize a = normalize a" by simp with * have "unit_factor (normalize a) * normalize a div normalize a = normalize a div normalize a" by simp with * show ?thesis by simp qed lemma normalize_dvd_iff [simp]: "normalize a dvd b \ a dvd b" proof - have "normalize a dvd b \ unit_factor a * normalize a dvd b" using mult_unit_dvd_iff [of "unit_factor a" "normalize a" b] by (cases "a = 0") simp_all then show ?thesis by simp qed lemma dvd_normalize_iff [simp]: "a dvd normalize b \ a dvd b" proof - have "a dvd normalize b \ a dvd normalize b * unit_factor b" using dvd_mult_unit_iff [of "unit_factor b" a "normalize b"] by (cases "b = 0") simp_all then show ?thesis by simp qed lemma normalize_idem_imp_unit_factor_eq: assumes "normalize a = a" shows "unit_factor a = of_bool (a \ 0)" proof (cases "a = 0") case True then show ?thesis by simp next case False then show ?thesis using assms unit_factor_normalize [of a] by simp qed lemma normalize_idem_imp_is_unit_iff: assumes "normalize a = a" shows "is_unit a \ a = 1" using assms by (cases "a = 0") (auto dest: is_unit_normalize) lemma coprime_normalize_left_iff [simp]: "coprime (normalize a) b \ coprime a b" by (rule; rule coprimeI) (auto intro: coprime_common_divisor) lemma coprime_normalize_right_iff [simp]: "coprime a (normalize b) \ coprime a b" using coprime_normalize_left_iff [of b a] by (simp add: ac_simps) text \ We avoid an explicit definition of associated elements but prefer explicit normalisation instead. In theory we could define an abbreviation like \<^prop>\associated a b \ normalize a = normalize b\ but this is counterproductive without suggestive infix syntax, which we do not want to sacrifice for this purpose here. \ lemma associatedI: assumes "a dvd b" and "b dvd a" shows "normalize a = normalize b" proof (cases "a = 0 \ b = 0") case True with assms show ?thesis by auto next case False from \a dvd b\ obtain c where b: "b = a * c" .. moreover from \b dvd a\ obtain d where a: "a = b * d" .. ultimately have "b * 1 = b * (c * d)" by (simp add: ac_simps) with False have "1 = c * d" unfolding mult_cancel_left by simp then have "is_unit c" and "is_unit d" by auto with a b show ?thesis by (simp add: is_unit_normalize) qed lemma associatedD1: "normalize a = normalize b \ a dvd b" using dvd_normalize_iff [of _ b, symmetric] normalize_dvd_iff [of a _, symmetric] by simp lemma associatedD2: "normalize a = normalize b \ b dvd a" using dvd_normalize_iff [of _ a, symmetric] normalize_dvd_iff [of b _, symmetric] by simp lemma associated_unit: "normalize a = normalize b \ is_unit a \ is_unit b" using dvd_unit_imp_unit by (auto dest!: associatedD1 associatedD2) lemma associated_iff_dvd: "normalize a = normalize b \ a dvd b \ b dvd a" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by (auto intro!: associatedI) next assume ?lhs then have "unit_factor a * normalize a = unit_factor a * normalize b" by simp then have *: "normalize b * unit_factor a = a" by (simp add: ac_simps) show ?rhs proof (cases "a = 0 \ b = 0") case True with \?lhs\ show ?thesis by auto next case False then have "b dvd normalize b * unit_factor a" and "normalize b * unit_factor a dvd b" by (simp_all add: mult_unit_dvd_iff dvd_mult_unit_iff) with * show ?thesis by simp qed qed lemma associated_eqI: assumes "a dvd b" and "b dvd a" assumes "normalize a = a" and "normalize b = b" shows "a = b" proof - from assms have "normalize a = normalize b" unfolding associated_iff_dvd by simp with \normalize a = a\ have "a = normalize b" by simp with \normalize b = b\ show "a = b" by simp qed lemma normalize_unit_factor_eqI: assumes "normalize a = normalize b" and "unit_factor a = unit_factor b" shows "a = b" proof - from assms have "unit_factor a * normalize a = unit_factor b * normalize b" by simp then show ?thesis by simp qed lemma normalize_mult_normalize_left [simp]: "normalize (normalize a * b) = normalize (a * b)" by (rule associated_eqI) (auto intro!: mult_dvd_mono) lemma normalize_mult_normalize_right [simp]: "normalize (a * normalize b) = normalize (a * b)" by (rule associated_eqI) (auto intro!: mult_dvd_mono) end class normalization_semidom_multiplicative = normalization_semidom + assumes unit_factor_mult: "unit_factor (a * b) = unit_factor a * unit_factor b" begin lemma normalize_mult: "normalize (a * b) = normalize a * normalize b" proof (cases "a = 0 \ b = 0") case True then show ?thesis by auto next case False have "unit_factor (a * b) * normalize (a * b) = a * b" by (rule unit_factor_mult_normalize) then have "normalize (a * b) = a * b div unit_factor (a * b)" by simp also have "\ = a * b div unit_factor (b * a)" by (simp add: ac_simps) also have "\ = a * b div unit_factor b div unit_factor a" using False by (simp add: unit_factor_mult is_unit_div_mult2_eq [symmetric]) also have "\ = a * (b div unit_factor b) div unit_factor a" using False by (subst unit_div_mult_swap) simp_all also have "\ = normalize a * normalize b" using False by (simp add: mult.commute [of a] mult.commute [of "normalize a"] unit_div_mult_swap [symmetric]) finally show ?thesis . qed lemma dvd_unit_factor_div: assumes "b dvd a" shows "unit_factor (a div b) = unit_factor a div unit_factor b" proof - from assms have "a = a div b * b" by simp then have "unit_factor a = unit_factor (a div b * b)" by simp then show ?thesis by (cases "b = 0") (simp_all add: unit_factor_mult) qed lemma dvd_normalize_div: assumes "b dvd a" shows "normalize (a div b) = normalize a div normalize b" proof - from assms have "a = a div b * b" by simp then have "normalize a = normalize (a div b * b)" by simp then show ?thesis by (cases "b = 0") (simp_all add: normalize_mult) qed end text \Syntactic division remainder operator\ class modulo = dvd + divide + fixes modulo :: "'a \ 'a \ 'a" (infixl "mod" 70) text \Arbitrary quotient and remainder partitions\ class semiring_modulo = comm_semiring_1_cancel + divide + modulo + assumes div_mult_mod_eq: "a div b * b + a mod b = a" begin lemma mod_div_decomp: fixes a b obtains q r where "q = a div b" and "r = a mod b" and "a = q * b + r" proof - from div_mult_mod_eq have "a = a div b * b + a mod b" by simp moreover have "a div b = a div b" .. moreover have "a mod b = a mod b" .. note that ultimately show thesis by blast qed lemma mult_div_mod_eq: "b * (a div b) + a mod b = a" using div_mult_mod_eq [of a b] by (simp add: ac_simps) lemma mod_div_mult_eq: "a mod b + a div b * b = a" using div_mult_mod_eq [of a b] by (simp add: ac_simps) lemma mod_mult_div_eq: "a mod b + b * (a div b) = a" using div_mult_mod_eq [of a b] by (simp add: ac_simps) lemma minus_div_mult_eq_mod: "a - a div b * b = a mod b" by (rule add_implies_diff [symmetric]) (fact mod_div_mult_eq) lemma minus_mult_div_eq_mod: "a - b * (a div b) = a mod b" by (rule add_implies_diff [symmetric]) (fact mod_mult_div_eq) lemma minus_mod_eq_div_mult: "a - a mod b = a div b * b" by (rule add_implies_diff [symmetric]) (fact div_mult_mod_eq) lemma minus_mod_eq_mult_div: "a - a mod b = b * (a div b)" by (rule add_implies_diff [symmetric]) (fact mult_div_mod_eq) lemma mod_0_imp_dvd [dest!]: "b dvd a" if "a mod b = 0" proof - have "b dvd (a div b) * b" by simp also have "(a div b) * b = a" using div_mult_mod_eq [of a b] by (simp add: that) finally show ?thesis . qed lemma [nitpick_unfold]: "a mod b = a - a div b * b" by (fact minus_div_mult_eq_mod [symmetric]) end subsection \Quotient and remainder in integral domains\ class semidom_modulo = algebraic_semidom + semiring_modulo begin lemma mod_0 [simp]: "0 mod a = 0" using div_mult_mod_eq [of 0 a] by simp lemma mod_by_0 [simp]: "a mod 0 = a" using div_mult_mod_eq [of a 0] by simp lemma mod_by_1 [simp]: "a mod 1 = 0" proof - from div_mult_mod_eq [of a one] div_by_1 have "a + a mod 1 = a" by simp then have "a + a mod 1 = a + 0" by simp then show ?thesis by (rule add_left_imp_eq) qed lemma mod_self [simp]: "a mod a = 0" using div_mult_mod_eq [of a a] by simp lemma dvd_imp_mod_0 [simp]: "b mod a = 0" if "a dvd b" using that minus_div_mult_eq_mod [of b a] by simp lemma mod_eq_0_iff_dvd: "a mod b = 0 \ b dvd a" by (auto intro: mod_0_imp_dvd) lemma dvd_eq_mod_eq_0 [nitpick_unfold, code]: "a dvd b \ b mod a = 0" by (simp add: mod_eq_0_iff_dvd) lemma dvd_mod_iff: assumes "c dvd b" shows "c dvd a mod b \ c dvd a" proof - from assms have "(c dvd a mod b) \ (c dvd ((a div b) * b + a mod b))" by (simp add: dvd_add_right_iff) also have "(a div b) * b + a mod b = a" using div_mult_mod_eq [of a b] by simp finally show ?thesis . qed lemma dvd_mod_imp_dvd: assumes "c dvd a mod b" and "c dvd b" shows "c dvd a" using assms dvd_mod_iff [of c b a] by simp lemma dvd_minus_mod [simp]: "b dvd a - a mod b" by (simp add: minus_mod_eq_div_mult) lemma cancel_div_mod_rules: "((a div b) * b + a mod b) + c = a + c" "(b * (a div b) + a mod b) + c = a + c" by (simp_all add: div_mult_mod_eq mult_div_mod_eq) end class idom_modulo = idom + semidom_modulo begin subclass idom_divide .. lemma div_diff [simp]: "c dvd a \ c dvd b \ (a - b) div c = a div c - b div c" using div_add [of _ _ "- b"] by (simp add: dvd_neg_div) end subsection \Interlude: basic tool support for algebraic and arithmetic calculations\ named_theorems arith "arith facts -- only ground formulas" ML_file \Tools/arith_data.ML\ ML_file \~~/src/Provers/Arith/cancel_div_mod.ML\ ML \ structure Cancel_Div_Mod_Ring = Cancel_Div_Mod ( val div_name = \<^const_name>\divide\; val mod_name = \<^const_name>\modulo\; val mk_binop = HOLogic.mk_binop; val mk_sum = Arith_Data.mk_sum; val dest_sum = Arith_Data.dest_sum; val div_mod_eqs = map mk_meta_eq @{thms cancel_div_mod_rules}; val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac @{thms diff_conv_add_uminus add_0_left add_0_right ac_simps}) ) \ simproc_setup cancel_div_mod_int ("(a::'a::semidom_modulo) + b") = \K Cancel_Div_Mod_Ring.proc\ subsection \Ordered semirings and rings\ text \ The theory of partially ordered rings is taken from the books: \<^item> \<^emph>\Lattice Theory\ by Garret Birkhoff, American Mathematical Society, 1979 \<^item> \<^emph>\Partially Ordered Algebraic Systems\, Pergamon Press, 1963 Most of the used notions can also be looked up in \<^item> \<^url>\http://www.mathworld.com\ by Eric Weisstein et. al. \<^item> \<^emph>\Algebra I\ by van der Waerden, Springer \ class ordered_semiring = semiring + ordered_comm_monoid_add + assumes mult_left_mono: "a \ b \ 0 \ c \ c * a \ c * b" assumes mult_right_mono: "a \ b \ 0 \ c \ a * c \ b * c" begin lemma mult_mono: "a \ b \ c \ d \ 0 \ b \ 0 \ c \ a * c \ b * d" apply (erule (1) mult_right_mono [THEN order_trans]) apply (erule (1) mult_left_mono) done lemma mult_mono': "a \ b \ c \ d \ 0 \ a \ 0 \ c \ a * c \ b * d" by (rule mult_mono) (fast intro: order_trans)+ end lemma mono_mult: fixes a :: "'a::ordered_semiring" shows "a \ 0 \ mono ((*) a)" by (simp add: mono_def mult_left_mono) class ordered_semiring_0 = semiring_0 + ordered_semiring begin lemma mult_nonneg_nonneg [simp]: "0 \ a \ 0 \ b \ 0 \ a * b" using mult_left_mono [of 0 b a] by simp lemma mult_nonneg_nonpos: "0 \ a \ b \ 0 \ a * b \ 0" using mult_left_mono [of b 0 a] by simp lemma mult_nonpos_nonneg: "a \ 0 \ 0 \ b \ a * b \ 0" using mult_right_mono [of a 0 b] by simp text \Legacy -- use @{thm [source] mult_nonpos_nonneg}.\ lemma mult_nonneg_nonpos2: "0 \ a \ b \ 0 \ b * a \ 0" by (drule mult_right_mono [of b 0]) auto lemma split_mult_neg_le: "(0 \ a \ b \ 0) \ (a \ 0 \ 0 \ b) \ a * b \ 0" by (auto simp add: mult_nonneg_nonpos mult_nonneg_nonpos2) end class ordered_cancel_semiring = ordered_semiring + cancel_comm_monoid_add begin subclass semiring_0_cancel .. subclass ordered_semiring_0 .. end class linordered_semiring = ordered_semiring + linordered_cancel_ab_semigroup_add begin subclass ordered_cancel_semiring .. subclass ordered_cancel_comm_monoid_add .. subclass ordered_ab_semigroup_monoid_add_imp_le .. lemma mult_left_less_imp_less: "c * a < c * b \ 0 \ c \ a < b" by (force simp add: mult_left_mono not_le [symmetric]) lemma mult_right_less_imp_less: "a * c < b * c \ 0 \ c \ a < b" by (force simp add: mult_right_mono not_le [symmetric]) end class zero_less_one = order + zero + one + assumes zero_less_one [simp]: "0 < 1" begin subclass zero_neq_one by standard (simp add: less_imp_neq) lemma zero_le_one [simp]: \0 \ 1\ by (rule less_imp_le) simp end class linordered_semiring_1 = linordered_semiring + semiring_1 + zero_less_one begin lemma convex_bound_le: assumes "x \ a" "y \ a" "0 \ u" "0 \ v" "u + v = 1" shows "u * x + v * y \ a" proof- from assms have "u * x + v * y \ u * a + v * a" by (simp add: add_mono mult_left_mono) with assms show ?thesis unfolding distrib_right[symmetric] by simp qed end class linordered_semiring_strict = semiring + comm_monoid_add + linordered_cancel_ab_semigroup_add + assumes mult_strict_left_mono: "a < b \ 0 < c \ c * a < c * b" assumes mult_strict_right_mono: "a < b \ 0 < c \ a * c < b * c" begin subclass semiring_0_cancel .. subclass linordered_semiring proof fix a b c :: 'a assume *: "a \ b" "0 \ c" then show "c * a \ c * b" unfolding le_less using mult_strict_left_mono by (cases "c = 0") auto from * show "a * c \ b * c" unfolding le_less using mult_strict_right_mono by (cases "c = 0") auto qed lemma mult_left_le_imp_le: "c * a \ c * b \ 0 < c \ a \ b" by (auto simp add: mult_strict_left_mono _not_less [symmetric]) lemma mult_right_le_imp_le: "a * c \ b * c \ 0 < c \ a \ b" by (auto simp add: mult_strict_right_mono not_less [symmetric]) lemma mult_pos_pos[simp]: "0 < a \ 0 < b \ 0 < a * b" using mult_strict_left_mono [of 0 b a] by simp lemma mult_pos_neg: "0 < a \ b < 0 \ a * b < 0" using mult_strict_left_mono [of b 0 a] by simp lemma mult_neg_pos: "a < 0 \ 0 < b \ a * b < 0" using mult_strict_right_mono [of a 0 b] by simp text \Legacy -- use @{thm [source] mult_neg_pos}.\ lemma mult_pos_neg2: "0 < a \ b < 0 \ b * a < 0" by (drule mult_strict_right_mono [of b 0]) auto lemma zero_less_mult_pos: assumes "0 < a * b" "0 < a" shows "0 < b" proof (cases "b \ 0") case True then show ?thesis using assms by (auto simp: le_less dest: less_not_sym mult_pos_neg [of a b]) qed (auto simp add: le_less not_less) lemma zero_less_mult_pos2: assumes "0 < b * a" "0 < a" shows "0 < b" proof (cases "b \ 0") case True then show ?thesis using assms by (auto simp: le_less dest: less_not_sym mult_pos_neg2 [of a b]) qed (auto simp add: le_less not_less) text \Strict monotonicity in both arguments\ lemma mult_strict_mono: assumes "a < b" "c < d" "0 < b" "0 \ c" shows "a * c < b * d" proof (cases "c = 0") case True with assms show ?thesis by simp next case False with assms have "a*c < b*c" by (simp add: mult_strict_right_mono [OF \a < b\]) also have "\ < b*d" by (simp add: assms mult_strict_left_mono) finally show ?thesis . qed text \This weaker variant has more natural premises\ lemma mult_strict_mono': assumes "a < b" and "c < d" and "0 \ a" and "0 \ c" shows "a * c < b * d" by (rule mult_strict_mono) (insert assms, auto) lemma mult_less_le_imp_less: assumes "a < b" and "c \ d" and "0 \ a" and "0 < c" shows "a * c < b * d" proof - have "a * c < b * c" by (simp add: assms mult_strict_right_mono) also have "... \ b * d" by (intro mult_left_mono) (use assms in auto) finally show ?thesis . qed lemma mult_le_less_imp_less: assumes "a \ b" and "c < d" and "0 < a" and "0 \ c" shows "a * c < b * d" proof - have "a * c \ b * c" by (simp add: assms mult_right_mono) also have "... < b * d" by (intro mult_strict_left_mono) (use assms in auto) finally show ?thesis . qed end class linordered_semiring_1_strict = linordered_semiring_strict + semiring_1 + zero_less_one begin subclass linordered_semiring_1 .. lemma convex_bound_lt: assumes "x < a" "y < a" "0 \ u" "0 \ v" "u + v = 1" shows "u * x + v * y < a" proof - from assms have "u * x + v * y < u * a + v * a" by (cases "u = 0") (auto intro!: add_less_le_mono mult_strict_left_mono mult_left_mono) with assms show ?thesis unfolding distrib_right[symmetric] by simp qed end class ordered_comm_semiring = comm_semiring_0 + ordered_ab_semigroup_add + assumes comm_mult_left_mono: "a \ b \ 0 \ c \ c * a \ c * b" begin subclass ordered_semiring proof fix a b c :: 'a assume "a \ b" "0 \ c" then show "c * a \ c * b" by (rule comm_mult_left_mono) then show "a * c \ b * c" by (simp only: mult.commute) qed end class ordered_cancel_comm_semiring = ordered_comm_semiring + cancel_comm_monoid_add begin subclass comm_semiring_0_cancel .. subclass ordered_comm_semiring .. subclass ordered_cancel_semiring .. end class linordered_comm_semiring_strict = comm_semiring_0 + linordered_cancel_ab_semigroup_add + assumes comm_mult_strict_left_mono: "a < b \ 0 < c \ c * a < c * b" begin subclass linordered_semiring_strict proof fix a b c :: 'a assume "a < b" "0 < c" then show "c * a < c * b" by (rule comm_mult_strict_left_mono) then show "a * c < b * c" by (simp only: mult.commute) qed subclass ordered_cancel_comm_semiring proof fix a b c :: 'a assume "a \ b" "0 \ c" then show "c * a \ c * b" unfolding le_less using mult_strict_left_mono by (cases "c = 0") auto qed end class ordered_ring = ring + ordered_cancel_semiring begin subclass ordered_ab_group_add .. lemma less_add_iff1: "a * e + c < b * e + d \ (a - b) * e + c < d" by (simp add: algebra_simps) lemma less_add_iff2: "a * e + c < b * e + d \ c < (b - a) * e + d" by (simp add: algebra_simps) lemma le_add_iff1: "a * e + c \ b * e + d \ (a - b) * e + c \ d" by (simp add: algebra_simps) lemma le_add_iff2: "a * e + c \ b * e + d \ c \ (b - a) * e + d" by (simp add: algebra_simps) lemma mult_left_mono_neg: "b \ a \ c \ 0 \ c * a \ c * b" by (auto dest: mult_left_mono [of _ _ "- c"]) lemma mult_right_mono_neg: "b \ a \ c \ 0 \ a * c \ b * c" by (auto dest: mult_right_mono [of _ _ "- c"]) lemma mult_nonpos_nonpos: "a \ 0 \ b \ 0 \ 0 \ a * b" using mult_right_mono_neg [of a 0 b] by simp lemma split_mult_pos_le: "(0 \ a \ 0 \ b) \ (a \ 0 \ b \ 0) \ 0 \ a * b" by (auto simp add: mult_nonpos_nonpos) end class abs_if = minus + uminus + ord + zero + abs + assumes abs_if: "\a\ = (if a < 0 then - a else a)" class linordered_ring = ring + linordered_semiring + linordered_ab_group_add + abs_if begin subclass ordered_ring .. subclass ordered_ab_group_add_abs proof fix a b show "\a + b\ \ \a\ + \b\" by (auto simp add: abs_if not_le not_less algebra_simps simp del: add.commute dest: add_neg_neg add_nonneg_nonneg) qed (auto simp: abs_if) lemma zero_le_square [simp]: "0 \ a * a" using linear [of 0 a] by (auto simp add: mult_nonpos_nonpos) lemma not_square_less_zero [simp]: "\ (a * a < 0)" by (simp add: not_less) proposition abs_eq_iff: "\x\ = \y\ \ x = y \ x = -y" by (auto simp add: abs_if split: if_split_asm) lemma abs_eq_iff': "\a\ = b \ b \ 0 \ (a = b \ a = - b)" by (cases "a \ 0") auto lemma eq_abs_iff': "a = \b\ \ a \ 0 \ (b = a \ b = - a)" using abs_eq_iff' [of b a] by auto lemma sum_squares_ge_zero: "0 \ x * x + y * y" by (intro add_nonneg_nonneg zero_le_square) lemma not_sum_squares_lt_zero: "\ x * x + y * y < 0" by (simp add: not_less sum_squares_ge_zero) end class linordered_ring_strict = ring + linordered_semiring_strict + ordered_ab_group_add + abs_if begin subclass linordered_ring .. lemma mult_strict_left_mono_neg: "b < a \ c < 0 \ c * a < c * b" using mult_strict_left_mono [of b a "- c"] by simp lemma mult_strict_right_mono_neg: "b < a \ c < 0 \ a * c < b * c" using mult_strict_right_mono [of b a "- c"] by simp lemma mult_neg_neg: "a < 0 \ b < 0 \ 0 < a * b" using mult_strict_right_mono_neg [of a 0 b] by simp subclass ring_no_zero_divisors proof fix a b assume "a \ 0" then have a: "a < 0 \ 0 < a" by (simp add: neq_iff) assume "b \ 0" then have b: "b < 0 \ 0 < b" by (simp add: neq_iff) have "a * b < 0 \ 0 < a * b" proof (cases "a < 0") case True show ?thesis proof (cases "b < 0") case True with \a < 0\ show ?thesis by (auto dest: mult_neg_neg) next case False with b have "0 < b" by auto with \a < 0\ show ?thesis by (auto dest: mult_strict_right_mono) qed next case False with a have "0 < a" by auto show ?thesis proof (cases "b < 0") case True with \0 < a\ show ?thesis by (auto dest: mult_strict_right_mono_neg) next case False with b have "0 < b" by auto with \0 < a\ show ?thesis by auto qed qed then show "a * b \ 0" by (simp add: neq_iff) qed lemma zero_less_mult_iff [algebra_split_simps, field_split_simps]: "0 < a * b \ 0 < a \ 0 < b \ a < 0 \ b < 0" by (cases a 0 b 0 rule: linorder_cases[case_product linorder_cases]) (auto simp add: mult_neg_neg not_less le_less dest: zero_less_mult_pos zero_less_mult_pos2) lemma zero_le_mult_iff [algebra_split_simps, field_split_simps]: "0 \ a * b \ 0 \ a \ 0 \ b \ a \ 0 \ b \ 0" by (auto simp add: eq_commute [of 0] le_less not_less zero_less_mult_iff) lemma mult_less_0_iff [algebra_split_simps, field_split_simps]: "a * b < 0 \ 0 < a \ b < 0 \ a < 0 \ 0 < b" using zero_less_mult_iff [of "- a" b] by auto lemma mult_le_0_iff [algebra_split_simps, field_split_simps]: "a * b \ 0 \ 0 \ a \ b \ 0 \ a \ 0 \ 0 \ b" using zero_le_mult_iff [of "- a" b] by auto text \ Cancellation laws for \<^term>\c * a < c * b\ and \<^term>\a * c < b * c\, also with the relations \\\ and equality. \ text \ These ``disjunction'' versions produce two cases when the comparison is an assumption, but effectively four when the comparison is a goal. \ lemma mult_less_cancel_right_disj: "a * c < b * c \ 0 < c \ a < b \ c < 0 \ b < a" proof (cases "c = 0") case False show ?thesis (is "?lhs \ ?rhs") proof assume ?lhs then have "c < 0 \ b < a" "c > 0 \ b > a" by (auto simp flip: not_le intro: mult_right_mono mult_right_mono_neg) with False show ?rhs by (auto simp add: neq_iff) next assume ?rhs with False show ?lhs by (auto simp add: mult_strict_right_mono mult_strict_right_mono_neg) qed qed auto lemma mult_less_cancel_left_disj: "c * a < c * b \ 0 < c \ a < b \ c < 0 \ b < a" proof (cases "c = 0") case False show ?thesis (is "?lhs \ ?rhs") proof assume ?lhs then have "c < 0 \ b < a" "c > 0 \ b > a" by (auto simp flip: not_le intro: mult_left_mono mult_left_mono_neg) with False show ?rhs by (auto simp add: neq_iff) next assume ?rhs with False show ?lhs by (auto simp add: mult_strict_left_mono mult_strict_left_mono_neg) qed qed auto text \ The ``conjunction of implication'' lemmas produce two cases when the comparison is a goal, but give four when the comparison is an assumption. \ lemma mult_less_cancel_right: "a * c < b * c \ (0 \ c \ a < b) \ (c \ 0 \ b < a)" using mult_less_cancel_right_disj [of a c b] by auto lemma mult_less_cancel_left: "c * a < c * b \ (0 \ c \ a < b) \ (c \ 0 \ b < a)" using mult_less_cancel_left_disj [of c a b] by auto lemma mult_le_cancel_right: "a * c \ b * c \ (0 < c \ a \ b) \ (c < 0 \ b \ a)" by (simp add: not_less [symmetric] mult_less_cancel_right_disj) lemma mult_le_cancel_left: "c * a \ c * b \ (0 < c \ a \ b) \ (c < 0 \ b \ a)" by (simp add: not_less [symmetric] mult_less_cancel_left_disj) lemma mult_le_cancel_left_pos: "0 < c \ c * a \ c * b \ a \ b" by (auto simp: mult_le_cancel_left) lemma mult_le_cancel_left_neg: "c < 0 \ c * a \ c * b \ b \ a" by (auto simp: mult_le_cancel_left) lemma mult_less_cancel_left_pos: "0 < c \ c * a < c * b \ a < b" by (auto simp: mult_less_cancel_left) lemma mult_less_cancel_left_neg: "c < 0 \ c * a < c * b \ b < a" by (auto simp: mult_less_cancel_left) end lemmas mult_sign_intros = mult_nonneg_nonneg mult_nonneg_nonpos mult_nonpos_nonneg mult_nonpos_nonpos mult_pos_pos mult_pos_neg mult_neg_pos mult_neg_neg class ordered_comm_ring = comm_ring + ordered_comm_semiring begin subclass ordered_ring .. subclass ordered_cancel_comm_semiring .. end class linordered_nonzero_semiring = ordered_comm_semiring + monoid_mult + linorder + zero_less_one + assumes add_mono1: "a < b \ a + 1 < b + 1" begin subclass zero_neq_one by standard (insert zero_less_one, blast) subclass comm_semiring_1 by standard (rule mult_1_left) lemma zero_le_one [simp]: "0 \ 1" by (rule zero_less_one [THEN less_imp_le]) lemma not_one_le_zero [simp]: "\ 1 \ 0" by (simp add: not_le) lemma not_one_less_zero [simp]: "\ 1 < 0" by (simp add: not_less) lemma of_bool_less_eq_iff [simp]: \of_bool P \ of_bool Q \ (P \ Q)\ by auto lemma of_bool_less_iff [simp]: \of_bool P < of_bool Q \ \ P \ Q\ by auto lemma mult_left_le: "c \ 1 \ 0 \ a \ a * c \ a" using mult_left_mono[of c 1 a] by simp lemma mult_le_one: "a \ 1 \ 0 \ b \ b \ 1 \ a * b \ 1" using mult_mono[of a 1 b 1] by simp lemma zero_less_two: "0 < 1 + 1" using add_pos_pos[OF zero_less_one zero_less_one] . end class linordered_semidom = semidom + linordered_comm_semiring_strict + zero_less_one + assumes le_add_diff_inverse2 [simp]: "b \ a \ a - b + b = a" begin subclass linordered_nonzero_semiring proof show "a + 1 < b + 1" if "a < b" for a b proof (rule ccontr, simp add: not_less) assume "b \ a" with that show False - by (simp add: ) + by simp qed qed lemma zero_less_eq_of_bool [simp]: \0 \ of_bool P\ by simp lemma zero_less_of_bool_iff [simp]: \0 < of_bool P \ P\ by simp lemma of_bool_less_eq_one [simp]: \of_bool P \ 1\ by simp lemma of_bool_less_one_iff [simp]: \of_bool P < 1 \ \ P\ by simp lemma of_bool_or_iff [simp]: \of_bool (P \ Q) = max (of_bool P) (of_bool Q)\ by (simp add: max_def) text \Addition is the inverse of subtraction.\ lemma le_add_diff_inverse [simp]: "b \ a \ b + (a - b) = a" by (frule le_add_diff_inverse2) (simp add: add.commute) lemma add_diff_inverse: "\ a < b \ b + (a - b) = a" by simp lemma add_le_imp_le_diff: assumes "i + k \ n" shows "i \ n - k" proof - have "n - (i + k) + i + k = n" by (simp add: assms add.assoc) with assms add_implies_diff have "i + k \ n - k + k" by fastforce then show ?thesis by simp qed lemma add_le_add_imp_diff_le: assumes 1: "i + k \ n" and 2: "n \ j + k" shows "i + k \ n \ n \ j + k \ n - k \ j" proof - have "n - (i + k) + i + k = n" using 1 by (simp add: add.assoc) moreover have "n - k = n - k - i + i" using 1 by (simp add: add_le_imp_le_diff) ultimately show ?thesis using 2 add_le_imp_le_diff [of "n-k" k "j + k"] by (simp add: add.commute diff_diff_add) qed lemma less_1_mult: "1 < m \ 1 < n \ 1 < m * n" using mult_strict_mono [of 1 m 1 n] by (simp add: less_trans [OF zero_less_one]) end class linordered_idom = comm_ring_1 + linordered_comm_semiring_strict + ordered_ab_group_add + abs_if + sgn + assumes sgn_if: "sgn x = (if x = 0 then 0 else if 0 < x then 1 else - 1)" begin subclass linordered_ring_strict .. subclass linordered_semiring_1_strict proof have "0 \ 1 * 1" by (fact zero_le_square) then show "0 < 1" by (simp add: le_less) qed subclass ordered_comm_ring .. subclass idom .. subclass linordered_semidom by standard simp subclass idom_abs_sgn by standard (auto simp add: sgn_if abs_if zero_less_mult_iff) lemma abs_bool_eq [simp]: \\of_bool P\ = of_bool P\ by simp lemma linorder_neqE_linordered_idom: assumes "x \ y" obtains "x < y" | "y < x" using assms by (rule neqE) text \These cancellation simp rules also produce two cases when the comparison is a goal.\ lemma mult_le_cancel_right1: "c \ b * c \ (0 < c \ 1 \ b) \ (c < 0 \ b \ 1)" using mult_le_cancel_right [of 1 c b] by simp lemma mult_le_cancel_right2: "a * c \ c \ (0 < c \ a \ 1) \ (c < 0 \ 1 \ a)" using mult_le_cancel_right [of a c 1] by simp lemma mult_le_cancel_left1: "c \ c * b \ (0 < c \ 1 \ b) \ (c < 0 \ b \ 1)" using mult_le_cancel_left [of c 1 b] by simp lemma mult_le_cancel_left2: "c * a \ c \ (0 < c \ a \ 1) \ (c < 0 \ 1 \ a)" using mult_le_cancel_left [of c a 1] by simp lemma mult_less_cancel_right1: "c < b * c \ (0 \ c \ 1 < b) \ (c \ 0 \ b < 1)" using mult_less_cancel_right [of 1 c b] by simp lemma mult_less_cancel_right2: "a * c < c \ (0 \ c \ a < 1) \ (c \ 0 \ 1 < a)" using mult_less_cancel_right [of a c 1] by simp lemma mult_less_cancel_left1: "c < c * b \ (0 \ c \ 1 < b) \ (c \ 0 \ b < 1)" using mult_less_cancel_left [of c 1 b] by simp lemma mult_less_cancel_left2: "c * a < c \ (0 \ c \ a < 1) \ (c \ 0 \ 1 < a)" using mult_less_cancel_left [of c a 1] by simp lemma sgn_0_0: "sgn a = 0 \ a = 0" by (fact sgn_eq_0_iff) lemma sgn_1_pos: "sgn a = 1 \ a > 0" unfolding sgn_if by simp lemma sgn_1_neg: "sgn a = - 1 \ a < 0" unfolding sgn_if by auto lemma sgn_pos [simp]: "0 < a \ sgn a = 1" by (simp only: sgn_1_pos) lemma sgn_neg [simp]: "a < 0 \ sgn a = - 1" by (simp only: sgn_1_neg) lemma abs_sgn: "\k\ = k * sgn k" unfolding sgn_if abs_if by auto lemma sgn_greater [simp]: "0 < sgn a \ 0 < a" unfolding sgn_if by auto lemma sgn_less [simp]: "sgn a < 0 \ a < 0" unfolding sgn_if by auto lemma abs_sgn_eq_1 [simp]: "a \ 0 \ \sgn a\ = 1" by simp lemma abs_sgn_eq: "\sgn a\ = (if a = 0 then 0 else 1)" by (simp add: sgn_if) lemma sgn_mult_self_eq [simp]: "sgn a * sgn a = of_bool (a \ 0)" by (cases "a > 0") simp_all lemma abs_mult_self_eq [simp]: "\a\ * \a\ = a * a" by (cases "a > 0") simp_all lemma same_sgn_sgn_add: "sgn (a + b) = sgn a" if "sgn b = sgn a" proof (cases a 0 rule: linorder_cases) case equal with that show ?thesis by simp next case less with that have "b < 0" by (simp add: sgn_1_neg) with \a < 0\ have "a + b < 0" by (rule add_neg_neg) with \a < 0\ show ?thesis by simp next case greater with that have "b > 0" by (simp add: sgn_1_pos) with \a > 0\ have "a + b > 0" by (rule add_pos_pos) with \a > 0\ show ?thesis by simp qed lemma same_sgn_abs_add: "\a + b\ = \a\ + \b\" if "sgn b = sgn a" proof - have "a + b = sgn a * \a\ + sgn b * \b\" by (simp add: sgn_mult_abs) also have "\ = sgn a * (\a\ + \b\)" using that by (simp add: algebra_simps) finally show ?thesis by (auto simp add: abs_mult) qed lemma sgn_not_eq_imp: "sgn a = - sgn b" if "sgn b \ sgn a" and "sgn a \ 0" and "sgn b \ 0" using that by (cases "a < 0") (auto simp add: sgn_0_0 sgn_1_pos sgn_1_neg) lemma abs_dvd_iff [simp]: "\m\ dvd k \ m dvd k" by (simp add: abs_if) lemma dvd_abs_iff [simp]: "m dvd \k\ \ m dvd k" by (simp add: abs_if) lemma dvd_if_abs_eq: "\l\ = \k\ \ l dvd k" by (subst abs_dvd_iff [symmetric]) simp text \ The following lemmas can be proven in more general structures, but are dangerous as simp rules in absence of @{thm neg_equal_zero}, @{thm neg_less_pos}, @{thm neg_less_eq_nonneg}. \ lemma equation_minus_iff_1 [simp, no_atp]: "1 = - a \ a = - 1" by (fact equation_minus_iff) lemma minus_equation_iff_1 [simp, no_atp]: "- a = 1 \ a = - 1" by (subst minus_equation_iff, auto) lemma le_minus_iff_1 [simp, no_atp]: "1 \ - b \ b \ - 1" by (fact le_minus_iff) lemma minus_le_iff_1 [simp, no_atp]: "- a \ 1 \ - 1 \ a" by (fact minus_le_iff) lemma less_minus_iff_1 [simp, no_atp]: "1 < - b \ b < - 1" by (fact less_minus_iff) lemma minus_less_iff_1 [simp, no_atp]: "- a < 1 \ - 1 < a" by (fact minus_less_iff) lemma add_less_zeroD: shows "x+y < 0 \ x<0 \ y<0" by (auto simp: not_less intro: le_less_trans [of _ "x+y"]) end text \Reasoning about inequalities with division\ context linordered_semidom begin lemma less_add_one: "a < a + 1" proof - have "a + 0 < a + 1" by (blast intro: zero_less_one add_strict_left_mono) then show ?thesis by simp qed end context linordered_idom begin lemma mult_right_le_one_le: "0 \ x \ 0 \ y \ y \ 1 \ x * y \ x" by (rule mult_left_le) lemma mult_left_le_one_le: "0 \ x \ 0 \ y \ y \ 1 \ y * x \ x" by (auto simp add: mult_le_cancel_right2) end text \Absolute Value\ context linordered_idom begin lemma mult_sgn_abs: "sgn x * \x\ = x" by (fact sgn_mult_abs) lemma abs_one: "\1\ = 1" by (fact abs_1) end class ordered_ring_abs = ordered_ring + ordered_ab_group_add_abs + assumes abs_eq_mult: "(0 \ a \ a \ 0) \ (0 \ b \ b \ 0) \ \a * b\ = \a\ * \b\" context linordered_idom begin subclass ordered_ring_abs by standard (auto simp: abs_if not_less mult_less_0_iff) lemma abs_mult_self: "\a\ * \a\ = a * a" by (fact abs_mult_self_eq) lemma abs_mult_less: assumes ac: "\a\ < c" and bd: "\b\ < d" shows "\a\ * \b\ < c * d" proof - from ac have "0 < c" by (blast intro: le_less_trans abs_ge_zero) with bd show ?thesis by (simp add: ac mult_strict_mono) qed lemma abs_less_iff: "\a\ < b \ a < b \ - a < b" by (simp add: less_le abs_le_iff) (auto simp add: abs_if) lemma abs_mult_pos: "0 \ x \ \y\ * x = \y * x\" by (simp add: abs_mult) lemma abs_diff_less_iff: "\x - a\ < r \ a - r < x \ x < a + r" by (auto simp add: diff_less_eq ac_simps abs_less_iff) lemma abs_diff_le_iff: "\x - a\ \ r \ a - r \ x \ x \ a + r" by (auto simp add: diff_le_eq ac_simps abs_le_iff) lemma abs_add_one_gt_zero: "0 < 1 + \x\" by (auto simp: abs_if not_less intro: zero_less_one add_strict_increasing less_trans) end subsection \Dioids\ text \ Dioids are the alternative extensions of semirings, a semiring can either be a ring or a dioid but never both. \ class dioid = semiring_1 + canonically_ordered_monoid_add begin subclass ordered_semiring by standard (auto simp: le_iff_add distrib_left distrib_right) end hide_fact (open) comm_mult_left_mono comm_mult_strict_left_mono distrib code_identifier code_module Rings \ (SML) Arith and (OCaml) Arith and (Haskell) Arith end diff --git a/src/HOL/Set_Interval.thy b/src/HOL/Set_Interval.thy --- a/src/HOL/Set_Interval.thy +++ b/src/HOL/Set_Interval.thy @@ -1,2578 +1,2578 @@ (* Title: HOL/Set_Interval.thy Author: Tobias Nipkow, Clemens Ballarin, Jeremy Avigad lessThan, greaterThan, atLeast, atMost and two-sided intervals Modern convention: Ixy stands for an interval where x and y describe the lower and upper bound and x,y : {c,o,i} where c = closed, o = open, i = infinite. Examples: Ico = {_ ..< _} and Ici = {_ ..} *) section \Set intervals\ theory Set_Interval imports Divides begin (* Belongs in Finite_Set but 2 is not available there *) lemma card_2_iff: "card S = 2 \ (\x y. S = {x,y} \ x \ y)" by (auto simp: card_Suc_eq numeral_eq_Suc) lemma card_2_iff': "card S = 2 \ (\x\S. \y\S. x \ y \ (\z\S. z = x \ z = y))" by (auto simp: card_Suc_eq numeral_eq_Suc) lemma card_3_iff: "card S = 3 \ (\x y z. S = {x,y,z} \ x \ y \ y \ z \ x \ z)" by (fastforce simp: card_Suc_eq numeral_eq_Suc) context ord begin definition lessThan :: "'a => 'a set" ("(1{..<_})") where "{.. 'a set" ("(1{.._})") where "{..u} == {x. x \ u}" definition greaterThan :: "'a => 'a set" ("(1{_<..})") where "{l<..} == {x. l 'a set" ("(1{_..})") where "{l..} == {x. l\x}" definition greaterThanLessThan :: "'a => 'a => 'a set" ("(1{_<..<_})") where "{l<.. 'a => 'a set" ("(1{_..<_})") where "{l.. 'a => 'a set" ("(1{_<.._})") where "{l<..u} == {l<..} Int {..u}" definition atLeastAtMost :: "'a => 'a => 'a set" ("(1{_.._})") where "{l..u} == {l..} Int {..u}" end text\A note of warning when using \<^term>\{.. on type \<^typ>\nat\: it is equivalent to \<^term>\{0::nat.. but some lemmas involving \<^term>\{m.. may not exist in \<^term>\{..-form as well.\ syntax (ASCII) "_UNION_le" :: "'a => 'a => 'b set => 'b set" ("(3UN _<=_./ _)" [0, 0, 10] 10) "_UNION_less" :: "'a => 'a => 'b set => 'b set" ("(3UN _<_./ _)" [0, 0, 10] 10) "_INTER_le" :: "'a => 'a => 'b set => 'b set" ("(3INT _<=_./ _)" [0, 0, 10] 10) "_INTER_less" :: "'a => 'a => 'b set => 'b set" ("(3INT _<_./ _)" [0, 0, 10] 10) syntax (latex output) "_UNION_le" :: "'a \ 'a => 'b set => 'b set" ("(3\(\unbreakable\_ \ _)/ _)" [0, 0, 10] 10) "_UNION_less" :: "'a \ 'a => 'b set => 'b set" ("(3\(\unbreakable\_ < _)/ _)" [0, 0, 10] 10) "_INTER_le" :: "'a \ 'a => 'b set => 'b set" ("(3\(\unbreakable\_ \ _)/ _)" [0, 0, 10] 10) "_INTER_less" :: "'a \ 'a => 'b set => 'b set" ("(3\(\unbreakable\_ < _)/ _)" [0, 0, 10] 10) syntax "_UNION_le" :: "'a => 'a => 'b set => 'b set" ("(3\_\_./ _)" [0, 0, 10] 10) "_UNION_less" :: "'a => 'a => 'b set => 'b set" ("(3\_<_./ _)" [0, 0, 10] 10) "_INTER_le" :: "'a => 'a => 'b set => 'b set" ("(3\_\_./ _)" [0, 0, 10] 10) "_INTER_less" :: "'a => 'a => 'b set => 'b set" ("(3\_<_./ _)" [0, 0, 10] 10) translations "\i\n. A" \ "\i\{..n}. A" "\i "\i\{..i\n. A" \ "\i\{..n}. A" "\i "\i\{..Various equivalences\ lemma (in ord) lessThan_iff [iff]: "(i \ lessThan k) = (i greaterThan k) = (k atLeast k) = (k<=i)" by (simp add: atLeast_def) lemma Compl_atLeast [simp]: "!!k:: 'a::linorder. -atLeast k = lessThan k" by (auto simp add: lessThan_def atLeast_def) lemma (in ord) atMost_iff [iff]: "(i \ atMost k) = (i<=k)" by (simp add: atMost_def) lemma atMost_Int_atLeast: "!!n:: 'a::order. atMost n Int atLeast n = {n}" by (blast intro: order_antisym) lemma (in linorder) lessThan_Int_lessThan: "{ a <..} \ { b <..} = { max a b <..}" by auto lemma (in linorder) greaterThan_Int_greaterThan: "{..< a} \ {..< b} = {..< min a b}" by auto subsection \Logical Equivalences for Set Inclusion and Equality\ lemma atLeast_empty_triv [simp]: "{{}..} = UNIV" by auto lemma atMost_UNIV_triv [simp]: "{..UNIV} = UNIV" by auto lemma atLeast_subset_iff [iff]: "(atLeast x \ atLeast y) = (y \ (x::'a::preorder))" by (blast intro: order_trans) lemma atLeast_eq_iff [iff]: "(atLeast x = atLeast y) = (x = (y::'a::order))" by (blast intro: order_antisym order_trans) lemma greaterThan_subset_iff [iff]: "(greaterThan x \ greaterThan y) = (y \ (x::'a::linorder))" unfolding greaterThan_def by (auto simp: linorder_not_less [symmetric]) lemma greaterThan_eq_iff [iff]: "(greaterThan x = greaterThan y) = (x = (y::'a::linorder))" by (auto simp: elim!: equalityE) lemma atMost_subset_iff [iff]: "(atMost x \ atMost y) = (x \ (y::'a::preorder))" by (blast intro: order_trans) lemma atMost_eq_iff [iff]: "(atMost x = atMost y) = (x = (y::'a::order))" by (blast intro: order_antisym order_trans) lemma lessThan_subset_iff [iff]: "(lessThan x \ lessThan y) = (x \ (y::'a::linorder))" unfolding lessThan_def by (auto simp: linorder_not_less [symmetric]) lemma lessThan_eq_iff [iff]: "(lessThan x = lessThan y) = (x = (y::'a::linorder))" by (auto simp: elim!: equalityE) lemma lessThan_strict_subset_iff: fixes m n :: "'a::linorder" shows "{.. m < n" by (metis leD lessThan_subset_iff linorder_linear not_less_iff_gr_or_eq psubset_eq) lemma (in linorder) Ici_subset_Ioi_iff: "{a ..} \ {b <..} \ b < a" by auto lemma (in linorder) Iic_subset_Iio_iff: "{.. a} \ {..< b} \ a < b" by auto lemma (in preorder) Ioi_le_Ico: "{a <..} \ {a ..}" by (auto intro: less_imp_le) subsection \Two-sided intervals\ context ord begin lemma greaterThanLessThan_iff [simp]: "(i \ {l<.. i < u)" by (simp add: greaterThanLessThan_def) lemma atLeastLessThan_iff [simp]: "(i \ {l.. i \ i < u)" by (simp add: atLeastLessThan_def) lemma greaterThanAtMost_iff [simp]: "(i \ {l<..u}) = (l < i \ i \ u)" by (simp add: greaterThanAtMost_def) lemma atLeastAtMost_iff [simp]: "(i \ {l..u}) = (l \ i \ i \ u)" by (simp add: atLeastAtMost_def) text \The above four lemmas could be declared as iffs. Unfortunately this breaks many proofs. Since it only helps blast, it is better to leave them alone.\ lemma greaterThanLessThan_eq: "{ a <..< b} = { a <..} \ {..< b }" by auto lemma (in order) atLeastLessThan_eq_atLeastAtMost_diff: "{a..Emptyness, singletons, subset\ context preorder begin lemma atLeastatMost_empty_iff[simp]: "{a..b} = {} \ (\ a \ b)" by auto (blast intro: order_trans) lemma atLeastatMost_empty_iff2[simp]: "{} = {a..b} \ (\ a \ b)" by auto (blast intro: order_trans) lemma atLeastLessThan_empty_iff[simp]: "{a.. (\ a < b)" by auto (blast intro: le_less_trans) lemma atLeastLessThan_empty_iff2[simp]: "{} = {a.. (\ a < b)" by auto (blast intro: le_less_trans) lemma greaterThanAtMost_empty_iff[simp]: "{k<..l} = {} \ \ k < l" by auto (blast intro: less_le_trans) lemma greaterThanAtMost_empty_iff2[simp]: "{} = {k<..l} \ \ k < l" by auto (blast intro: less_le_trans) lemma atLeastatMost_subset_iff[simp]: "{a..b} \ {c..d} \ (\ a \ b) \ c \ a \ b \ d" unfolding atLeastAtMost_def atLeast_def atMost_def by (blast intro: order_trans) lemma atLeastatMost_psubset_iff: "{a..b} < {c..d} \ ((\ a \ b) \ c \ a \ b \ d \ (c < a \ b < d)) \ c \ d" by(simp add: psubset_eq set_eq_iff less_le_not_le)(blast intro: order_trans) lemma atLeastAtMost_subseteq_atLeastLessThan_iff: "{a..b} \ {c ..< d} \ (a \ b \ c \ a \ b < d)" by auto (blast intro: local.order_trans local.le_less_trans elim: )+ lemma Icc_subset_Ici_iff[simp]: "{l..h} \ {l'..} = (\ l\h \ l\l')" by(auto simp: subset_eq intro: order_trans) lemma Icc_subset_Iic_iff[simp]: "{l..h} \ {..h'} = (\ l\h \ h\h')" by(auto simp: subset_eq intro: order_trans) lemma not_Ici_eq_empty[simp]: "{l..} \ {}" by(auto simp: set_eq_iff) lemma not_Iic_eq_empty[simp]: "{..h} \ {}" by(auto simp: set_eq_iff) lemmas not_empty_eq_Ici_eq_empty[simp] = not_Ici_eq_empty[symmetric] lemmas not_empty_eq_Iic_eq_empty[simp] = not_Iic_eq_empty[symmetric] end context order begin lemma atLeastatMost_empty[simp]: "b < a \ {a..b} = {}" by(auto simp: atLeastAtMost_def atLeast_def atMost_def) lemma atLeastLessThan_empty[simp]: "b \ a \ {a.. k ==> {k<..l} = {}" by(auto simp:greaterThanAtMost_def greaterThan_def atMost_def) lemma greaterThanLessThan_empty[simp]:"l \ k ==> {k<.. {a .. b} = {a}" by simp lemma Icc_eq_Icc[simp]: "{l..h} = {l'..h'} = (l=l' \ h=h' \ \ l\h \ \ l'\h')" by (simp add: order_class.order.eq_iff) (auto intro: order_trans) lemma atLeastAtMost_singleton_iff[simp]: "{a .. b} = {c} \ a = b \ b = c" proof assume "{a..b} = {c}" hence *: "\ (\ a \ b)" unfolding atLeastatMost_empty_iff[symmetric] by simp with \{a..b} = {c}\ have "c \ a \ b \ c" by auto with * show "a = b \ b = c" by auto qed simp end context no_top begin (* also holds for no_bot but no_top should suffice *) lemma not_UNIV_le_Icc[simp]: "\ UNIV \ {l..h}" using gt_ex[of h] by(auto simp: subset_eq less_le_not_le) lemma not_UNIV_le_Iic[simp]: "\ UNIV \ {..h}" using gt_ex[of h] by(auto simp: subset_eq less_le_not_le) lemma not_Ici_le_Icc[simp]: "\ {l..} \ {l'..h'}" using gt_ex[of h'] by(auto simp: subset_eq less_le)(blast dest:antisym_conv intro: order_trans) lemma not_Ici_le_Iic[simp]: "\ {l..} \ {..h'}" using gt_ex[of h'] by(auto simp: subset_eq less_le)(blast dest:antisym_conv intro: order_trans) end context no_bot begin lemma not_UNIV_le_Ici[simp]: "\ UNIV \ {l..}" using lt_ex[of l] by(auto simp: subset_eq less_le_not_le) lemma not_Iic_le_Icc[simp]: "\ {..h} \ {l'..h'}" using lt_ex[of l'] by(auto simp: subset_eq less_le)(blast dest:antisym_conv intro: order_trans) lemma not_Iic_le_Ici[simp]: "\ {..h} \ {l'..}" using lt_ex[of l'] by(auto simp: subset_eq less_le)(blast dest:antisym_conv intro: order_trans) end context no_top begin (* also holds for no_bot but no_top should suffice *) lemma not_UNIV_eq_Icc[simp]: "\ UNIV = {l'..h'}" using gt_ex[of h'] by(auto simp: set_eq_iff less_le_not_le) lemmas not_Icc_eq_UNIV[simp] = not_UNIV_eq_Icc[symmetric] lemma not_UNIV_eq_Iic[simp]: "\ UNIV = {..h'}" using gt_ex[of h'] by(auto simp: set_eq_iff less_le_not_le) lemmas not_Iic_eq_UNIV[simp] = not_UNIV_eq_Iic[symmetric] lemma not_Icc_eq_Ici[simp]: "\ {l..h} = {l'..}" unfolding atLeastAtMost_def using not_Ici_le_Iic[of l'] by blast lemmas not_Ici_eq_Icc[simp] = not_Icc_eq_Ici[symmetric] (* also holds for no_bot but no_top should suffice *) lemma not_Iic_eq_Ici[simp]: "\ {..h} = {l'..}" using not_Ici_le_Iic[of l' h] by blast lemmas not_Ici_eq_Iic[simp] = not_Iic_eq_Ici[symmetric] end context no_bot begin lemma not_UNIV_eq_Ici[simp]: "\ UNIV = {l'..}" using lt_ex[of l'] by(auto simp: set_eq_iff less_le_not_le) lemmas not_Ici_eq_UNIV[simp] = not_UNIV_eq_Ici[symmetric] lemma not_Icc_eq_Iic[simp]: "\ {l..h} = {..h'}" unfolding atLeastAtMost_def using not_Iic_le_Ici[of h'] by blast lemmas not_Iic_eq_Icc[simp] = not_Icc_eq_Iic[symmetric] end context dense_linorder begin lemma greaterThanLessThan_empty_iff[simp]: "{ a <..< b } = {} \ b \ a" using dense[of a b] by (cases "a < b") auto lemma greaterThanLessThan_empty_iff2[simp]: "{} = { a <..< b } \ b \ a" using dense[of a b] by (cases "a < b") auto lemma atLeastLessThan_subseteq_atLeastAtMost_iff: "{a ..< b} \ { c .. d } \ (a < b \ c \ a \ b \ d)" using dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanAtMost_subseteq_atLeastAtMost_iff: "{a <.. b} \ { c .. d } \ (a < b \ c \ a \ b \ d)" using dense[of "a" "min c b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanLessThan_subseteq_atLeastAtMost_iff: "{a <..< b} \ { c .. d } \ (a < b \ c \ a \ b \ d)" using dense[of "a" "min c b"] dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanLessThan_subseteq_greaterThanLessThan: "{a <..< b} \ {c <..< d} \ (a < b \ a \ c \ b \ d)" using dense[of "a" "min c b"] dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanAtMost_subseteq_atLeastLessThan_iff: "{a <.. b} \ { c ..< d } \ (a < b \ c \ a \ b < d)" using dense[of "a" "min c b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanLessThan_subseteq_atLeastLessThan_iff: "{a <..< b} \ { c ..< d } \ (a < b \ c \ a \ b \ d)" using dense[of "a" "min c b"] dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) lemma greaterThanLessThan_subseteq_greaterThanAtMost_iff: "{a <..< b} \ { c <.. d } \ (a < b \ c \ a \ b \ d)" using dense[of "a" "min c b"] dense[of "max a d" "b"] by (force simp: subset_eq Ball_def not_less[symmetric]) end context no_top begin lemma greaterThan_non_empty[simp]: "{x <..} \ {}" using gt_ex[of x] by auto end context no_bot begin lemma lessThan_non_empty[simp]: "{..< x} \ {}" using lt_ex[of x] by auto end lemma (in linorder) atLeastLessThan_subset_iff: "{a.. {c.. b \ a \ c\a \ b\d" apply (auto simp:subset_eq Ball_def not_le) apply(frule_tac x=a in spec) apply(erule_tac x=d in allE) - apply (auto simp: ) + apply auto done lemma atLeastLessThan_inj: fixes a b c d :: "'a::linorder" assumes eq: "{a ..< b} = {c ..< d}" and "a < b" "c < d" shows "a = c" "b = d" using assms by (metis atLeastLessThan_subset_iff eq less_le_not_le antisym_conv2 subset_refl)+ lemma atLeastLessThan_eq_iff: fixes a b c d :: "'a::linorder" assumes "a < b" "c < d" shows "{a ..< b} = {c ..< d} \ a = c \ b = d" using atLeastLessThan_inj assms by auto lemma (in linorder) Ioc_inj: \{a <.. b} = {c <.. d} \ (b \ a \ d \ c) \ a = c \ b = d\ (is \?P \ ?Q\) proof assume ?Q then show ?P by auto next assume ?P then have \a < x \ x \ b \ c < x \ x \ d\ for x by (simp add: set_eq_iff) from this [of a] this [of b] this [of c] this [of d] show ?Q by auto qed lemma (in order) Iio_Int_singleton: "{.. {x} = (if x < k then {x} else {})" by auto lemma (in linorder) Ioc_subset_iff: "{a<..b} \ {c<..d} \ (b \ a \ c \ a \ b \ d)" by (auto simp: subset_eq Ball_def) (metis less_le not_less) lemma (in order_bot) atLeast_eq_UNIV_iff: "{x..} = UNIV \ x = bot" by (auto simp: set_eq_iff intro: le_bot) lemma (in order_top) atMost_eq_UNIV_iff: "{..x} = UNIV \ x = top" by (auto simp: set_eq_iff intro: top_le) lemma (in bounded_lattice) atLeastAtMost_eq_UNIV_iff: "{x..y} = UNIV \ (x = bot \ y = top)" by (auto simp: set_eq_iff intro: top_le le_bot) lemma Iio_eq_empty_iff: "{..< n::'a::{linorder, order_bot}} = {} \ n = bot" by (auto simp: set_eq_iff not_less le_bot) lemma lessThan_empty_iff: "{..< n::nat} = {} \ n = 0" by (simp add: Iio_eq_empty_iff bot_nat_def) lemma mono_image_least: assumes f_mono: "mono f" and f_img: "f ` {m ..< n} = {m' ..< n'}" "m < n" shows "f m = m'" proof - from f_img have "{m' ..< n'} \ {}" by (metis atLeastLessThan_empty_iff image_is_empty) with f_img have "m' \ f ` {m ..< n}" by auto then obtain k where "f k = m'" "m \ k" by auto moreover have "m' \ f m" using f_img by auto ultimately show "f m = m'" using f_mono by (auto elim: monoE[where x=m and y=k]) qed subsection \Infinite intervals\ context dense_linorder begin lemma infinite_Ioo: assumes "a < b" shows "\ finite {a<.. {}" using \a < b\ by auto ultimately have "a < Max {a <..< b}" "Max {a <..< b} < b" using Max_in[of "{a <..< b}"] by auto then obtain x where "Max {a <..< b} < x" "x < b" using dense[of "Max {a<.. {a <..< b}" using \a < Max {a <..< b}\ by auto then have "x \ Max {a <..< b}" using fin by auto with \Max {a <..< b} < x\ show False by auto qed lemma infinite_Icc: "a < b \ \ finite {a .. b}" using greaterThanLessThan_subseteq_atLeastAtMost_iff[of a b a b] infinite_Ioo[of a b] by (auto dest: finite_subset) lemma infinite_Ico: "a < b \ \ finite {a ..< b}" using greaterThanLessThan_subseteq_atLeastLessThan_iff[of a b a b] infinite_Ioo[of a b] by (auto dest: finite_subset) lemma infinite_Ioc: "a < b \ \ finite {a <.. b}" using greaterThanLessThan_subseteq_greaterThanAtMost_iff[of a b a b] infinite_Ioo[of a b] by (auto dest: finite_subset) lemma infinite_Ioo_iff [simp]: "infinite {a<.. a < b" using not_less_iff_gr_or_eq by (fastforce simp: infinite_Ioo) lemma infinite_Icc_iff [simp]: "infinite {a .. b} \ a < b" using not_less_iff_gr_or_eq by (fastforce simp: infinite_Icc) lemma infinite_Ico_iff [simp]: "infinite {a.. a < b" using not_less_iff_gr_or_eq by (fastforce simp: infinite_Ico) lemma infinite_Ioc_iff [simp]: "infinite {a<..b} \ a < b" using not_less_iff_gr_or_eq by (fastforce simp: infinite_Ioc) end lemma infinite_Iio: "\ finite {..< a :: 'a :: {no_bot, linorder}}" proof assume "finite {..< a}" then have *: "\x. x < a \ Min {..< a} \ x" by auto obtain x where "x < a" using lt_ex by auto obtain y where "y < Min {..< a}" using lt_ex by auto also have "Min {..< a} \ x" using \x < a\ by fact also note \x < a\ finally have "Min {..< a} \ y" by fact with \y < Min {..< a}\ show False by auto qed lemma infinite_Iic: "\ finite {.. a :: 'a :: {no_bot, linorder}}" using infinite_Iio[of a] finite_subset[of "{..< a}" "{.. a}"] by (auto simp: subset_eq less_imp_le) lemma infinite_Ioi: "\ finite {a :: 'a :: {no_top, linorder} <..}" proof assume "finite {a <..}" then have *: "\x. a < x \ x \ Max {a <..}" by auto obtain y where "Max {a <..} < y" using gt_ex by auto obtain x where x: "a < x" using gt_ex by auto also from x have "x \ Max {a <..}" by fact also note \Max {a <..} < y\ finally have "y \ Max { a <..}" by fact with \Max {a <..} < y\ show False by auto qed lemma infinite_Ici: "\ finite {a :: 'a :: {no_top, linorder} ..}" using infinite_Ioi[of a] finite_subset[of "{a <..}" "{a ..}"] by (auto simp: subset_eq less_imp_le) subsubsection \Intersection\ context linorder begin lemma Int_atLeastAtMost[simp]: "{a..b} Int {c..d} = {max a c .. min b d}" by auto lemma Int_atLeastAtMostR1[simp]: "{..b} Int {c..d} = {c .. min b d}" by auto lemma Int_atLeastAtMostR2[simp]: "{a..} Int {c..d} = {max a c .. d}" by auto lemma Int_atLeastAtMostL1[simp]: "{a..b} Int {..d} = {a .. min b d}" by auto lemma Int_atLeastAtMostL2[simp]: "{a..b} Int {c..} = {max a c .. b}" by auto lemma Int_atLeastLessThan[simp]: "{a.. {..b} = {.. min a b}" by (auto simp: min_def) lemma Ioc_disjoint: "{a<..b} \ {c<..d} = {} \ b \ a \ d \ c \ b \ c \ d \ a" by auto end context complete_lattice begin lemma shows Sup_atLeast[simp]: "Sup {x ..} = top" and Sup_greaterThanAtLeast[simp]: "x < top \ Sup {x <..} = top" and Sup_atMost[simp]: "Sup {.. y} = y" and Sup_atLeastAtMost[simp]: "x \ y \ Sup { x .. y} = y" and Sup_greaterThanAtMost[simp]: "x < y \ Sup { x <.. y} = y" by (auto intro!: Sup_eqI) lemma shows Inf_atMost[simp]: "Inf {.. x} = bot" and Inf_atMostLessThan[simp]: "top < x \ Inf {..< x} = bot" and Inf_atLeast[simp]: "Inf {x ..} = x" and Inf_atLeastAtMost[simp]: "x \ y \ Inf { x .. y} = x" and Inf_atLeastLessThan[simp]: "x < y \ Inf { x ..< y} = x" by (auto intro!: Inf_eqI) end lemma fixes x y :: "'a :: {complete_lattice, dense_linorder}" shows Sup_lessThan[simp]: "Sup {..< y} = y" and Sup_atLeastLessThan[simp]: "x < y \ Sup { x ..< y} = y" and Sup_greaterThanLessThan[simp]: "x < y \ Sup { x <..< y} = y" and Inf_greaterThan[simp]: "Inf {x <..} = x" and Inf_greaterThanAtMost[simp]: "x < y \ Inf { x <.. y} = x" and Inf_greaterThanLessThan[simp]: "x < y \ Inf { x <..< y} = x" by (auto intro!: Inf_eqI Sup_eqI intro: dense_le dense_le_bounded dense_ge dense_ge_bounded) subsection \Intervals of natural numbers\ subsubsection \The Constant \<^term>\lessThan\\ lemma lessThan_0 [simp]: "lessThan (0::nat) = {}" by (simp add: lessThan_def) lemma lessThan_Suc: "lessThan (Suc k) = insert k (lessThan k)" by (simp add: lessThan_def less_Suc_eq, blast) text \The following proof is convenient in induction proofs where new elements get indices at the beginning. So it is used to transform \<^term>\{.. to \<^term>\0::nat\ and \<^term>\{..< n}\.\ lemma zero_notin_Suc_image [simp]: "0 \ Suc ` A" by auto lemma lessThan_Suc_eq_insert_0: "{..m::nat. lessThan m) = UNIV" by blast subsubsection \The Constant \<^term>\greaterThan\\ lemma greaterThan_0: "greaterThan 0 = range Suc" unfolding greaterThan_def by (blast dest: gr0_conv_Suc [THEN iffD1]) lemma greaterThan_Suc: "greaterThan (Suc k) = greaterThan k - {Suc k}" unfolding greaterThan_def by (auto elim: linorder_neqE) lemma INT_greaterThan_UNIV: "(\m::nat. greaterThan m) = {}" by blast subsubsection \The Constant \<^term>\atLeast\\ lemma atLeast_0 [simp]: "atLeast (0::nat) = UNIV" by (unfold atLeast_def UNIV_def, simp) lemma atLeast_Suc: "atLeast (Suc k) = atLeast k - {k}" unfolding atLeast_def by (auto simp: order_le_less Suc_le_eq) lemma atLeast_Suc_greaterThan: "atLeast (Suc k) = greaterThan k" by (auto simp add: greaterThan_def atLeast_def less_Suc_eq_le) lemma UN_atLeast_UNIV: "(\m::nat. atLeast m) = UNIV" by blast subsubsection \The Constant \<^term>\atMost\\ lemma atMost_0 [simp]: "atMost (0::nat) = {0}" by (simp add: atMost_def) lemma atMost_Suc: "atMost (Suc k) = insert (Suc k) (atMost k)" unfolding atMost_def by (auto simp add: less_Suc_eq order_le_less) lemma UN_atMost_UNIV: "(\m::nat. atMost m) = UNIV" by blast subsubsection \The Constant \<^term>\atLeastLessThan\\ text\The orientation of the following 2 rules is tricky. The lhs is defined in terms of the rhs. Hence the chosen orientation makes sense in this theory --- the reverse orientation complicates proofs (eg nontermination). But outside, when the definition of the lhs is rarely used, the opposite orientation seems preferable because it reduces a specific concept to a more general one.\ lemma atLeast0LessThan [code_abbrev]: "{0::nat..The Constant \<^term>\atLeastAtMost\\ lemma Icc_eq_insert_lb_nat: "m \ n \ {m..n} = insert m {Suc m..n}" by auto lemma atLeast0_atMost_Suc: "{0..Suc n} = insert (Suc n) {0..n}" by (simp add: atLeast0AtMost atMost_Suc) lemma atLeast0_atMost_Suc_eq_insert_0: "{0..Suc n} = insert 0 (Suc ` {0..n})" by (simp add: atLeast0AtMost atMost_Suc_eq_insert_0) subsubsection \Intervals of nats with \<^term>\Suc\\ text\Not a simprule because the RHS is too messy.\ lemma atLeastLessThanSuc: "{m.. n then insert n {m.. Suc n \ {m..Suc n} = insert (Suc n) {m..n}" by auto lemma atLeastAtMost_insertL: "m \ n \ insert m {Suc m..n} = {m ..n}" by auto text \The analogous result is useful on \<^typ>\int\:\ (* here, because we don't have an own int section *) lemma atLeastAtMostPlus1_int_conv: "m \ 1+n \ {m..1+n} = insert (1+n) {m..n::int}" by (auto intro: set_eqI) lemma atLeastLessThan_add_Un: "i \ j \ {i.. {j..Intervals and numerals\ lemma lessThan_nat_numeral: \ \Evaluation for specific numerals\ "lessThan (numeral k :: nat) = insert (pred_numeral k) (lessThan (pred_numeral k))" by (simp add: numeral_eq_Suc lessThan_Suc) lemma atMost_nat_numeral: \ \Evaluation for specific numerals\ "atMost (numeral k :: nat) = insert (numeral k) (atMost (pred_numeral k))" by (simp add: numeral_eq_Suc atMost_Suc) lemma atLeastLessThan_nat_numeral: \ \Evaluation for specific numerals\ "atLeastLessThan m (numeral k :: nat) = (if m \ (pred_numeral k) then insert (pred_numeral k) (atLeastLessThan m (pred_numeral k)) else {})" by (simp add: numeral_eq_Suc atLeastLessThanSuc) subsubsection \Image\ context linordered_semidom begin lemma image_add_atLeast[simp]: "plus k ` {i..} = {k + i..}" proof - have "n = k + (n - k)" if "i + k \ n" for n proof - have "n = (n - (k + i)) + (k + i)" using that by (metis add_commute le_add_diff_inverse) then show "n = k + (n - k)" by (metis local.add_diff_cancel_left' add_assoc add_commute) qed then show ?thesis by (fastforce simp: add_le_imp_le_diff add.commute) qed lemma image_add_atLeastAtMost [simp]: "plus k ` {i..j} = {i + k..j + k}" (is "?A = ?B") proof show "?A \ ?B" by (auto simp add: ac_simps) next show "?B \ ?A" proof fix n assume "n \ ?B" then have "i \ n - k" by (simp add: add_le_imp_le_diff) have "n = n - k + k" proof - from \n \ ?B\ have "n = n - (i + k) + (i + k)" by simp also have "\ = n - k - i + i + k" by (simp add: algebra_simps) also have "\ = n - k + k" using \i \ n - k\ by simp finally show ?thesis . qed moreover have "n - k \ {i..j}" using \n \ ?B\ by (auto simp: add_le_imp_le_diff add_le_add_imp_diff_le) ultimately show "n \ ?A" by (simp add: ac_simps) qed qed lemma image_add_atLeastAtMost' [simp]: "(\n. n + k) ` {i..j} = {i + k..j + k}" by (simp add: add.commute [of _ k]) lemma image_add_atLeastLessThan [simp]: "plus k ` {i..n. n + k) ` {i.. uminus ` {x<..}" by (rule imageI) (simp add: *) thus "y \ uminus ` {x<..}" by simp next fix y assume "y \ -x" have "- (-y) \ uminus ` {x..}" by (rule imageI) (insert \y \ -x\[THEN le_imp_neg_le], simp) thus "y \ uminus ` {x..}" by simp qed simp_all lemma fixes x :: 'a shows image_uminus_lessThan[simp]: "uminus ` {.. = {c - b<..c - a}" by simp finally show ?thesis by simp qed lemma image_minus_const_greaterThanAtMost[simp]: fixes a b c::"'a::linordered_idom" shows "(-) c ` {a<..b} = {c - b.. = {c - b.. = {..c - a}" by simp finally show ?thesis by simp qed lemma image_minus_const_AtMost[simp]: fixes b c::"'a::linordered_idom" shows "(-) c ` {..b} = {c - b..}" proof - have "(-) c ` {..b} = (+) c ` uminus ` {..b}" unfolding image_image by simp also have "\ = {c - b..}" by simp finally show ?thesis by simp qed lemma image_minus_const_atLeastAtMost' [simp]: "(\t. t-d)`{a..b} = {a-d..b-d}" for d::"'a::linordered_idom" by (metis (no_types, lifting) diff_conv_add_uminus image_add_atLeastAtMost' image_cong) context linordered_field begin lemma image_mult_atLeastAtMost [simp]: "((*) d ` {a..b}) = {d*a..d*b}" if "d>0" using that by (auto simp: field_simps mult_le_cancel_right intro: rev_image_eqI [where x="x/d" for x]) lemma image_divide_atLeastAtMost [simp]: "((\c. c / d) ` {a..b}) = {a/d..b/d}" if "d>0" proof - from that have "inverse d > 0" by simp with image_mult_atLeastAtMost [of "inverse d" a b] have "(*) (inverse d) ` {a..b} = {inverse d * a..inverse d * b}" by blast moreover have "(*) (inverse d) = (\c. c / d)" by (simp add: fun_eq_iff field_simps) ultimately show ?thesis by simp qed lemma image_mult_atLeastAtMost_if: "(*) c ` {x .. y} = (if c > 0 then {c * x .. c * y} else if x \ y then {c * y .. c * x} else {})" proof (cases "c = 0 \ x > y") case True then show ?thesis by auto next case False then have "x \ y" by auto from False consider "c < 0"| "c > 0" by (auto simp add: neq_iff) then show ?thesis proof cases case 1 have "(*) c ` {x..y} = {c * y..c * x}" proof (rule set_eqI) fix d from 1 have "inj (\z. z / c)" by (auto intro: injI) then have "d \ (*) c ` {x..y} \ d / c \ (\z. z div c) ` (*) c ` {x..y}" by (subst inj_image_mem_iff) simp_all also have "\ \ d / c \ {x..y}" using 1 by (simp add: image_image) also have "\ \ d \ {c * y..c * x}" by (auto simp add: field_simps 1) finally show "d \ (*) c ` {x..y} \ d \ {c * y..c * x}" . qed with \x \ y\ show ?thesis by auto qed (simp add: mult_left_mono_neg) qed lemma image_mult_atLeastAtMost_if': "(\x. x * c) ` {x..y} = (if x \ y then if c > 0 then {x * c .. y * c} else {y * c .. x * c} else {})" using image_mult_atLeastAtMost_if [of c x y] by (auto simp add: ac_simps) lemma image_affinity_atLeastAtMost: "((\x. m * x + c) ` {a..b}) = (if {a..b} = {} then {} else if 0 \ m then {m * a + c .. m * b + c} else {m * b + c .. m * a + c})" proof - have *: "(\x. m * x + c) = ((\x. x + c) \ (*) m)" by (simp add: fun_eq_iff) show ?thesis by (simp only: * image_comp [symmetric] image_mult_atLeastAtMost_if) (auto simp add: mult_le_cancel_left) qed lemma image_affinity_atLeastAtMost_diff: "((\x. m*x - c) ` {a..b}) = (if {a..b}={} then {} else if 0 \ m then {m*a - c .. m*b - c} else {m*b - c .. m*a - c})" using image_affinity_atLeastAtMost [of m "-c" a b] by simp lemma image_affinity_atLeastAtMost_div: "((\x. x/m + c) ` {a..b}) = (if {a..b}={} then {} else if 0 \ m then {a/m + c .. b/m + c} else {b/m + c .. a/m + c})" using image_affinity_atLeastAtMost [of "inverse m" c a b] by (simp add: field_class.field_divide_inverse algebra_simps inverse_eq_divide) lemma image_affinity_atLeastAtMost_div_diff: "((\x. x/m - c) ` {a..b}) = (if {a..b}={} then {} else if 0 \ m then {a/m - c .. b/m - c} else {b/m - c .. a/m - c})" using image_affinity_atLeastAtMost_diff [of "inverse m" c a b] by (simp add: field_class.field_divide_inverse algebra_simps inverse_eq_divide) end lemma atLeast1_lessThan_eq_remove0: "{Suc 0..x. x + (l::int)) ` {0..i. i - c) ` {x ..< y} = (if c < y then {x - c ..< y - c} else if x < y then {0} else {})" (is "_ = ?right") proof safe fix a assume a: "a \ ?right" show "a \ (\i. i - c) ` {x ..< y}" proof cases assume "c < y" with a show ?thesis by (auto intro!: image_eqI[of _ _ "a + c"]) next assume "\ c < y" with a show ?thesis by (auto intro!: image_eqI[of _ _ x] split: if_split_asm) qed qed auto lemma image_int_atLeastLessThan: "int ` {a..Finiteness\ lemma finite_lessThan [iff]: fixes k :: nat shows "finite {..A bounded set of natural numbers is finite.\ lemma bounded_nat_set_is_finite: "(\i\N. i < (n::nat)) \ finite N" by (rule finite_subset [OF _ finite_lessThan]) auto text \A set of natural numbers is finite iff it is bounded.\ lemma finite_nat_set_iff_bounded: "finite(N::nat set) = (\m. \n\N. n?F\, simplified less_Suc_eq_le[symmetric]] by blast next assume ?B show ?F using \?B\ by(blast intro:bounded_nat_set_is_finite) qed lemma finite_nat_set_iff_bounded_le: "finite(N::nat set) = (\m. \n\N. n\m)" unfolding finite_nat_set_iff_bounded by (blast dest:less_imp_le_nat le_imp_less_Suc) lemma finite_less_ub: "!!f::nat=>nat. (!!n. n \ f n) ==> finite {n. f n \ u}" by (rule_tac B="{..u}" in finite_subset, auto intro: order_trans) lemma bounded_Max_nat: fixes P :: "nat \ bool" assumes x: "P x" and M: "\x. P x \ x \ M" obtains m where "P m" "\x. P x \ x \ m" proof - have "finite {x. P x}" using M finite_nat_set_iff_bounded_le by auto then have "Max {x. P x} \ {x. P x}" using Max_in x by auto then show ?thesis by (simp add: \finite {x. P x}\ that) qed text\Any subset of an interval of natural numbers the size of the subset is exactly that interval.\ lemma subset_card_intvl_is_intvl: assumes "A \ {k.. A" by auto with insert have "A \ {k..Proving Inclusions and Equalities between Unions\ lemma UN_le_eq_Un0: "(\i\n::nat. M i) = (\i\{1..n}. M i) \ M 0" (is "?A = ?B") proof show "?A \ ?B" proof fix x assume "x \ ?A" then obtain i where i: "i\n" "x \ M i" by auto show "x \ ?B" proof(cases i) case 0 with i show ?thesis by simp next case (Suc j) with i show ?thesis by auto qed qed next show "?B \ ?A" by fastforce qed lemma UN_le_add_shift: "(\i\n::nat. M(i+k)) = (\i\{k..n+k}. M i)" (is "?A = ?B") proof show "?A \ ?B" by fastforce next show "?B \ ?A" proof fix x assume "x \ ?B" then obtain i where i: "i \ {k..n+k}" "x \ M(i)" by auto hence "i-k\n \ x \ M((i-k)+k)" by auto thus "x \ ?A" by blast qed qed lemma UN_le_add_shift_strict: "(\ii\{k.. ?A" proof fix x assume "x \ ?B" then obtain i where i: "i \ {k.. M(i)" by auto then have "i - k < n \ x \ M((i-k) + k)" by auto then show "x \ ?A" using UN_le_add_shift by blast qed qed (fastforce) lemma UN_UN_finite_eq: "(\n::nat. \i\{0..n. A n)" by (auto simp add: atLeast0LessThan) lemma UN_finite_subset: "(\n::nat. (\i\{0.. C) \ (\n. A n) \ C" by (subst UN_UN_finite_eq [symmetric]) blast lemma UN_finite2_subset: assumes "\n::nat. (\i\{0.. (\i\{0..n. A n) \ (\n. B n)" proof (rule UN_finite_subset, rule) fix n and a from assms have "(\i\{0.. (\i\{0.. (\i\{0.. (\i\{0.. (\i. B i)" by (auto simp add: UN_UN_finite_eq) qed lemma UN_finite2_eq: "(\n::nat. (\i\{0..i\{0.. (\n. A n) = (\n. B n)" apply (rule subset_antisym [OF UN_finite_subset UN_finite2_subset]) apply auto apply (force simp add: atLeastLessThan_add_Un [of 0])+ done subsubsection \Cardinality\ lemma card_lessThan [simp]: "card {..x. x + l) ` {.. {0.. {0..n}" shows "finite N" using assms finite_atLeastAtMost by (rule finite_subset) lemma ex_bij_betw_nat_finite: "finite M \ \h. bij_betw h {0.. \h. bij_betw h M {0.. finite B \ card A = card B \ \h. bij_betw h A B" apply(drule ex_bij_betw_finite_nat) apply(drule ex_bij_betw_nat_finite) apply(auto intro!:bij_betw_trans) done lemma ex_bij_betw_nat_finite_1: "finite M \ \h. bij_betw h {1 .. card M} M" by (rule finite_same_card_bij) auto lemma bij_betw_iff_card: assumes "finite A" "finite B" shows "(\f. bij_betw f A B) \ (card A = card B)" proof assume "card A = card B" moreover obtain f where "bij_betw f A {0 ..< card A}" using assms ex_bij_betw_finite_nat by blast moreover obtain g where "bij_betw g {0 ..< card B} B" using assms ex_bij_betw_nat_finite by blast ultimately have "bij_betw (g \ f) A B" by (auto simp: bij_betw_trans) thus "(\f. bij_betw f A B)" by blast qed (auto simp: bij_betw_same_card) lemma subset_eq_atLeast0_lessThan_card: fixes n :: nat assumes "N \ {0.. n" proof - from assms finite_lessThan have "card N \ card {0..Relational version of @{thm [source] card_inj_on_le}:\ lemma card_le_if_inj_on_rel: assumes "finite B" "\a. a \ A \ \b. b\B \ r a b" "\a1 a2 b. \ a1 \ A; a2 \ A; b \ B; r a1 b; r a2 b \ \ a1 = a2" shows "card A \ card B" proof - let ?P = "\a b. b \ B \ r a b" let ?f = "\a. SOME b. ?P a b" have 1: "?f ` A \ B" by (auto intro: someI2_ex[OF assms(2)]) have "inj_on ?f A" proof (auto simp: inj_on_def) fix a1 a2 assume asms: "a1 \ A" "a2 \ A" "?f a1 = ?f a2" have 0: "?f a1 \ B" using "1" \a1 \ A\ by blast have 1: "r a1 (?f a1)" using someI_ex[OF assms(2)[OF \a1 \ A\]] by blast have 2: "r a2 (?f a1)" using someI_ex[OF assms(2)[OF \a2 \ A\]] asms(3) by auto show "a1 = a2" using assms(3)[OF asms(1,2) 0 1 2] . qed with 1 show ?thesis using card_inj_on_le[of ?f A B] assms(1) by simp qed lemma inj_on_funpow_least: \<^marker>\contributor \Lars Noschinski\\ \inj_on (\k. (f ^^ k) s) {0.. if \(f ^^ n) s = s\ \\m. 0 < m \ m < n \ (f ^^ m) s \ s\ proof - { fix k l assume A: "k < n" "l < n" "k \ l" "(f ^^ k) s = (f ^^ l) s" define k' l' where "k' = min k l" and "l' = max k l" with A have A': "k' < l'" "(f ^^ k') s = (f ^^ l') s" "l' < n" by (auto simp: min_def max_def) have "s = (f ^^ ((n - l') + l')) s" using that \l' < n\ by simp also have "\ = (f ^^ (n - l')) ((f ^^ l') s)" by (simp add: funpow_add) also have "(f ^^ l') s = (f ^^ k') s" by (simp add: A') also have "(f ^^ (n - l')) \ = (f ^^ (n - l' + k')) s" by (simp add: funpow_add) finally have "(f ^^ (n - l' + k')) s = s" by simp moreover have "n - l' + k' < n" "0 < n - l' + k'"using A' by linarith+ ultimately have False using that(2) by auto } then show ?thesis by (intro inj_onI) auto qed subsection \Intervals of integers\ lemma atLeastLessThanPlusOne_atLeastAtMost_int: "{l..Finiteness\ lemma image_atLeastZeroLessThan_int: "0 \ u ==> {(0::int).. u") case True then show ?thesis by (auto simp: image_atLeastZeroLessThan_int) qed auto lemma finite_atLeastLessThan_int [iff]: "finite {l..Cardinality\ lemma card_atLeastZeroLessThan_int: "card {(0::int).. u") case True then show ?thesis by (auto simp: image_atLeastZeroLessThan_int card_image inj_on_def) qed auto lemma card_atLeastLessThan_int [simp]: "card {l.. k < (i::nat)}" proof - have "{k. P k \ k < i} \ {.. M" shows "card {k \ M. k < Suc i} \ 0" proof - from zero_in_M have "{k \ M. k < Suc i} \ {}" by auto with finite_M_bounded_by_nat show ?thesis by (auto simp add: card_eq_0_iff) qed lemma card_less_Suc2: assumes "0 \ M" shows "card {k. Suc k \ M \ k < i} = card {k \ M. k < Suc i}" proof - have *: "\j \ M; j < Suc i\ \ j - Suc 0 < i \ Suc (j - Suc 0) \ M \ Suc 0 \ j" for j by (cases j) (use assms in auto) show ?thesis proof (rule card_bij_eq) show "inj_on Suc {k. Suc k \ M \ k < i}" by force show "inj_on (\x. x - Suc 0) {k \ M. k < Suc i}" by (rule inj_on_diff_nat) (use * in blast) qed (use * in auto) qed lemma card_less_Suc: assumes "0 \ M" shows "Suc (card {k. Suc k \ M \ k < i}) = card {k \ M. k < Suc i}" proof - have "Suc (card {k. Suc k \ M \ k < i}) = Suc (card {k. Suc k \ M - {0} \ k < i})" by simp also have "\ = Suc (card {k \ M - {0}. k < Suc i})" apply (subst card_less_Suc2) using assms by auto also have "\ = Suc (card ({k \ M. k < Suc i} - {0}))" by (force intro: arg_cong [where f=card]) also have "\ = card (insert 0 ({k \ M. k < Suc i} - {0}))" by (simp add: card.insert_remove) also have "... = card {k \ M. k < Suc i}" using assms by (force simp add: intro: arg_cong [where f=card]) finally show ?thesis. qed lemma card_le_Suc_Max: "finite S \ card S \ Suc (Max S)" proof (rule classical) assume "finite S" and "\ Suc (Max S) \ card S" then have "Suc (Max S) < card S" by simp with \finite S\ have "S \ {0..Max S}" by auto hence "card S \ card {0..Max S}" by (intro card_mono; auto) thus "card S \ Suc (Max S)" by simp qed subsection \Lemmas useful with the summation operator sum\ text \For examples, see Algebra/poly/UnivPoly2.thy\ subsubsection \Disjoint Unions\ text \Singletons and open intervals\ lemma ivl_disj_un_singleton: "{l::'a::linorder} Un {l<..} = {l..}" "{.. {l} Un {l<.. {l<.. u ==> {l} Un {l<..u} = {l..u}" "(l::'a::linorder) \ u ==> {l..One- and two-sided intervals\ lemma ivl_disj_un_one: "(l::'a::linorder) < u ==> {..l} Un {l<.. u ==> {.. u ==> {..l} Un {l<..u} = {..u}" "(l::'a::linorder) \ u ==> {.. u ==> {l<..u} Un {u<..} = {l<..}" "(l::'a::linorder) < u ==> {l<.. u ==> {l..u} Un {u<..} = {l..}" "(l::'a::linorder) \ u ==> {l..Two- and two-sided intervals\ lemma ivl_disj_un_two: "[| (l::'a::linorder) < m; m \ u |] ==> {l<.. m; m < u |] ==> {l<..m} Un {m<.. m; m \ u |] ==> {l.. m; m < u |] ==> {l..m} Un {m<.. u |] ==> {l<.. m; m \ u |] ==> {l<..m} Un {m<..u} = {l<..u}" "[| (l::'a::linorder) \ m; m \ u |] ==> {l.. m; m \ u |] ==> {l..m} Un {m<..u} = {l..u}" by auto lemma ivl_disj_un_two_touch: "[| (l::'a::linorder) < m; m < u |] ==> {l<..m} Un {m.. m; m < u |] ==> {l..m} Un {m.. u |] ==> {l<..m} Un {m..u} = {l<..u}" "[| (l::'a::linorder) \ m; m \ u |] ==> {l..m} Un {m..u} = {l..u}" by auto lemmas ivl_disj_un = ivl_disj_un_singleton ivl_disj_un_one ivl_disj_un_two ivl_disj_un_two_touch subsubsection \Disjoint Intersections\ text \One- and two-sided intervals\ lemma ivl_disj_int_one: "{..l::'a::order} Int {l<..Two- and two-sided intervals\ lemma ivl_disj_int_two: "{l::'a::order<..Some Differences\ lemma ivl_diff[simp]: "i \ n \ {i..Some Subset Conditions\ lemma ivl_subset [simp]: "({i.. {m.. i \ m \ i \ j \ (n::'a::linorder))" using linorder_class.le_less_linear[of i n] apply (auto simp: linorder_not_le) apply (force intro: leI)+ done lemma obtain_subset_with_card_n: assumes "n \ card S" obtains T where "T \ S" "card T = n" "finite T" proof - obtain n' where "card S = n + n'" by (metis assms le_add_diff_inverse) with that show thesis proof (induct n' arbitrary: S) case 0 then show ?case by (cases "finite S") auto next case Suc then show ?case by (simp add: card_Suc_eq) (metis subset_insertI2) qed qed subsection \Generic big monoid operation over intervals\ context semiring_char_0 begin lemma inj_on_of_nat [simp]: "inj_on of_nat N" by rule simp lemma bij_betw_of_nat [simp]: "bij_betw of_nat N A \ of_nat ` N = A" by (simp add: bij_betw_def) lemma Nats_infinite: "infinite (\ :: 'a set)" by (metis Nats_def finite_imageD infinite_UNIV_char_0 inj_on_of_nat) end context comm_monoid_set begin lemma atLeastLessThan_reindex: "F g {h m.. h) {m.. h) {m..n}" if "bij_betw h {m..n} {h m..h n}" for m n ::nat proof - from that have "inj_on h {m..n}" and "h ` {m..n} = {h m..h n}" by (simp_all add: bij_betw_def) then show ?thesis using reindex [of h "{m..n}" g] by simp qed lemma atLeastLessThan_shift_bounds: "F g {m + k.. plus k) {m.. plus k) {m..n}" for m n k :: nat using atLeastAtMost_reindex [of "plus k" m n g] by (simp add: ac_simps) lemma atLeast_Suc_lessThan_Suc_shift: "F g {Suc m.. Suc) {m.. Suc) {m..n}" using atLeastAtMost_shift_bounds [of _ _ 1] by (simp add: plus_1_eq_Suc) lemma atLeast_atMost_pred_shift: "F (g \ (\n. n - Suc 0)) {Suc m..Suc n} = F g {m..n}" unfolding atLeast_Suc_atMost_Suc_shift by simp lemma atLeast_lessThan_pred_shift: "F (g \ (\n. n - Suc 0)) {Suc m.. int) {m.. int) {m..n}" by (rule atLeastAtMost_reindex) (simp add: image_int_atLeastAtMost) lemma atLeast0_lessThan_Suc: "F g {0..* g n" by (simp add: atLeast0_lessThan_Suc ac_simps) lemma atLeast0_atMost_Suc: "F g {0..Suc n} = F g {0..n} \<^bold>* g (Suc n)" by (simp add: atLeast0_atMost_Suc ac_simps) lemma atLeast0_lessThan_Suc_shift: "F g {0..* F (g \ Suc) {0..* F (g \ Suc) {0..n}" by (simp add: atLeast0_atMost_Suc_eq_insert_0 atLeast_Suc_atMost_Suc_shift) lemma atLeast_Suc_lessThan: "F g {m..* F g {Suc m..* F g {Suc m..n}" if "m \ n" proof - from that have "{m..n} = insert m {Suc m..n}" by auto then show ?thesis by simp qed lemma ivl_cong: "a = c \ b = d \ (\x. c \ x \ x < d \ g x = h x) \ F g {a.. plus m) {0.. n") simp_all lemma atLeastAtMost_shift_0: fixes m n p :: nat assumes "m \ n" shows "F g {m..n} = F (g \ plus m) {0..n - m}" using assms atLeastAtMost_shift_bounds [of g 0 m "n - m"] by simp lemma atLeastLessThan_concat: fixes m n p :: nat shows "m \ n \ n \ p \ F g {m..* F g {n..i. g (m + n - Suc i)) {n..i. g (m + n - i)) {n..m}" by (rule reindex_bij_witness [where i="\i. m + n - i" and j="\i. m + n - i"]) auto lemma atLeastLessThan_rev_at_least_Suc_atMost: "F g {n..i. g (m + n - i)) {Suc n..m}" unfolding atLeastLessThan_rev [of g n m] by (cases m) (simp_all add: atLeast_Suc_atMost_Suc_shift atLeastLessThanSuc_atLeastAtMost) end subsection \Summation indexed over intervals\ syntax (ASCII) "_from_to_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(SUM _ = _.._./ _)" [0,0,0,10] 10) "_from_upto_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(SUM _ = _..<_./ _)" [0,0,0,10] 10) "_upt_sum" :: "idt \ 'a \ 'b \ 'b" ("(SUM _<_./ _)" [0,0,10] 10) "_upto_sum" :: "idt \ 'a \ 'b \ 'b" ("(SUM _<=_./ _)" [0,0,10] 10) syntax (latex_sum output) "_from_to_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\sum_{\_ = _\<^latex>\}^{\_\<^latex>\}$\ _)" [0,0,0,10] 10) "_from_upto_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\sum_{\_ = _\<^latex>\}^{<\_\<^latex>\}$\ _)" [0,0,0,10] 10) "_upt_sum" :: "idt \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\sum_{\_ < _\<^latex>\}$\ _)" [0,0,10] 10) "_upto_sum" :: "idt \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\sum_{\_ \ _\<^latex>\}$\ _)" [0,0,10] 10) syntax "_from_to_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _.._./ _)" [0,0,0,10] 10) "_from_upto_sum" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _..<_./ _)" [0,0,0,10] 10) "_upt_sum" :: "idt \ 'a \ 'b \ 'b" ("(3\_<_./ _)" [0,0,10] 10) "_upto_sum" :: "idt \ 'a \ 'b \ 'b" ("(3\_\_./ _)" [0,0,10] 10) translations "\x=a..b. t" == "CONST sum (\x. t) {a..b}" "\x=a..x. t) {a..i\n. t" == "CONST sum (\i. t) {..n}" "\ii. t) {..The above introduces some pretty alternative syntaxes for summation over intervals: \begin{center} \begin{tabular}{lll} Old & New & \LaTeX\\ @{term[source]"\x\{a..b}. e"} & \<^term>\\x=a..b. e\ & @{term[mode=latex_sum]"\x=a..b. e"}\\ @{term[source]"\x\{a..\\x=a.. & @{term[mode=latex_sum]"\x=a..x\{..b}. e"} & \<^term>\\x\b. e\ & @{term[mode=latex_sum]"\x\b. e"}\\ @{term[source]"\x\{..\\x & @{term[mode=latex_sum]"\xlatex_sum\ (e.g.\ via \mode = latex_sum\ in antiquotations). It is not the default \LaTeX\ output because it only works well with italic-style formulae, not tt-style. Note that for uniformity on \<^typ>\nat\ it is better to use \<^term>\\x::nat=0.. rather than \\x: \sum\ may not provide all lemmas available for \<^term>\{m.. also in the special form for \<^term>\{...\ text\This congruence rule should be used for sums over intervals as the standard theorem @{text[source]sum.cong} does not work well with the simplifier who adds the unsimplified premise \<^term>\x\B\ to the context.\ context comm_monoid_set begin lemma zero_middle: assumes "1 \ p" "k \ p" shows "F (\j. if j < k then g j else if j = k then \<^bold>1 else h (j - Suc 0)) {..p} = F (\j. if j < k then g j else h j) {..p - Suc 0}" (is "?lhs = ?rhs") proof - have [simp]: "{..p - Suc 0} \ {j. j < k} = {.. - {j. j < k} = {k..p - Suc 0}" using assms by auto have "?lhs = F g {..* F (\j. if j = k then \<^bold>1 else h (j - Suc 0)) {k..p}" using union_disjoint [of "{.. = F g {..* F (\j. h (j - Suc 0)) {Suc k..p}" by (simp add: atLeast_Suc_atMost [of k p] assms) also have "\ = F g {..* F h {k .. p - Suc 0}" using reindex [of Suc "{k..p - Suc 0}"] assms by simp also have "\ = ?rhs" by (simp add: If_cases) finally show ?thesis . qed lemma atMost_Suc [simp]: "F g {..Suc n} = F g {..n} \<^bold>* g (Suc n)" by (simp add: atMost_Suc ac_simps) lemma lessThan_Suc [simp]: "F g {..* g n" by (simp add: lessThan_Suc ac_simps) lemma cl_ivl_Suc [simp]: "F g {m..Suc n} = (if Suc n < m then \<^bold>1 else F g {m..n} \<^bold>* g(Suc n))" by (auto simp: ac_simps atLeastAtMostSuc_conv) lemma op_ivl_Suc [simp]: "F g {m..1 else F g {m..* g(n))" by (auto simp: ac_simps atLeastLessThanSuc) lemma head: fixes n :: nat assumes mn: "m \ n" shows "F g {m..n} = g m \<^bold>* F g {m<..n}" (is "?lhs = ?rhs") proof - from mn have "{m..n} = {m} \ {m<..n}" by (auto intro: ivl_disj_un_singleton) hence "?lhs = F g ({m} \ {m<..n})" by (simp add: atLeast0LessThan) also have "\ = ?rhs" by simp finally show ?thesis . qed lemma last_plus: fixes n::nat shows "m \ n \ F g {m..n} = g n \<^bold>* F g {m..1 else F g {m..* g(n))" by (simp add: commute last_plus) lemma ub_add_nat: assumes "(m::nat) \ n + 1" shows "F g {m..n + p} = F g {m..n} \<^bold>* F g {n + 1..n + p}" proof- have "{m .. n+p} = {m..n} \ {n+1..n+p}" using \m \ n+1\ by auto thus ?thesis by (auto simp: ivl_disj_int union_disjoint atLeastSucAtMost_greaterThanAtMost) qed lemma nat_group: fixes k::nat shows "F (\m. F g {m * k ..< m*k + k}) {.. 0" by auto then show ?thesis by (induct n) (simp_all add: atLeastLessThan_concat add.commute atLeast0LessThan[symmetric]) qed auto lemma triangle_reindex: fixes n :: nat shows "F (\(i,j). g i j) {(i,j). i+j < n} = F (\k. F (\i. g i (k - i)) {..k}) {..(i,j). g i j) {(i,j). i+j \ n} = F (\k. F (\i. g i (k - i)) {..k}) {..n}" using triangle_reindex [of g "Suc n"] by (simp only: Nat.less_Suc_eq_le lessThan_Suc_atMost) lemma nat_diff_reindex: "F (\i. g (n - Suc i)) {..i. g(i + k)){m..i. g(i + k)){m..n::nat}" by (rule reindex_bij_witness[where i="\i. i + k" and j="\i. i - k"]) auto corollary shift_bounds_cl_Suc_ivl: "F g {Suc m..Suc n} = F (\i. g(Suc i)){m..n}" by (simp add: shift_bounds_cl_nat_ivl[where k="Suc 0", simplified]) corollary Suc_reindex_ivl: "m \ n \ F g {m..n} \<^bold>* g (Suc n) = g m \<^bold>* F (\i. g (Suc i)) {m..n}" by (simp add: assoc atLeast_Suc_atMost flip: shift_bounds_cl_Suc_ivl) corollary shift_bounds_Suc_ivl: "F g {Suc m..i. g(Suc i)){m..* F (\i. g (Suc i)) {..n}" proof (induct n) case 0 show ?case by simp next case (Suc n) note IH = this have "F g {..Suc (Suc n)} = F g {..Suc n} \<^bold>* g (Suc (Suc n))" by (rule atMost_Suc) also have "F g {..Suc n} = g 0 \<^bold>* F (\i. g (Suc i)) {..n}" by (rule IH) also have "g 0 \<^bold>* F (\i. g (Suc i)) {..n} \<^bold>* g (Suc (Suc n)) = g 0 \<^bold>* (F (\i. g (Suc i)) {..n} \<^bold>* g (Suc (Suc n)))" by (rule assoc) also have "F (\i. g (Suc i)) {..n} \<^bold>* g (Suc (Suc n)) = F (\i. g (Suc i)) {..Suc n}" by (rule atMost_Suc [symmetric]) finally show ?case . qed lemma lessThan_Suc_shift: "F g {..* F (\i. g (Suc i)) {..* F (\i. g (Suc i)) {..i. F (\j. a i j) {0..j. F (\i. a i j) {Suc j..n}) {0..i. F (\j. a i j) {..j. F (\i. a i j) {Suc j..n}) {..k. g (Suc k)) {.. = F (\k. g (Suc k)) {.. b \ F g {a..* g b" by (simp add: atLeastLessThanSuc commute) lemma nat_ivl_Suc': assumes "m \ Suc n" shows "F g {m..Suc n} = g (Suc n) \<^bold>* F g {m..n}" proof - from assms have "{m..Suc n} = insert (Suc n) {m..n}" by auto also have "F g \ = g (Suc n) \<^bold>* F g {m..n}" by simp finally show ?thesis . qed lemma in_pairs: "F g {2*m..Suc(2*n)} = F (\i. g(2*i) \<^bold>* g(Suc(2*i))) {m..n}" proof (induction n) case 0 show ?case by (cases "m=0") auto next case (Suc n) then show ?case by (auto simp: assoc split: if_split_asm) qed lemma in_pairs_0: "F g {..Suc(2*n)} = F (\i. g(2*i) \<^bold>* g(Suc(2*i))) {..n}" using in_pairs [of _ 0 n] by (simp add: atLeast0AtMost) end lemma card_sum_le_nat_sum: "\ {0.. \ S" proof (cases "finite S") case True then show ?thesis proof (induction "card S" arbitrary: S) case (Suc x) then have "Max S \ x" using card_le_Suc_Max by fastforce let ?S' = "S - {Max S}" from Suc have "Max S \ S" by (auto intro: Max_in) hence cards: "card S = Suc (card ?S')" using \finite S\ by (intro card.remove; auto) hence "\ {0.. \ ?S'" using Suc by (intro Suc; auto) hence "\ {0.. \ ?S' + Max S" using \Max S \ x\ by simp also have "... = \ S" using sum.remove[OF \finite S\ \Max S \ S\, where g="\x. x"] by simp finally show ?case using cards Suc by auto qed simp qed simp lemma sum_natinterval_diff: fixes f:: "nat \ ('a::ab_group_add)" shows "sum (\k. f k - f(k + 1)) {(m::nat) .. n} = (if m \ n then f m - f(n + 1) else 0)" by (induct n, auto simp add: algebra_simps not_le le_Suc_eq) lemma sum_diff_nat_ivl: fixes f :: "nat \ 'a::ab_group_add" shows "\ m \ n; n \ p \ \ sum f {m..x. Q x \ P x \ (\xxxShifting bounds\ context comm_monoid_add begin context fixes f :: "nat \ 'a" assumes "f 0 = 0" begin lemma sum_shift_lb_Suc0_0_upt: "sum f {Suc 0..f 0 = 0\ by simp qed lemma sum_shift_lb_Suc0_0: "sum f {Suc 0..k} = sum f {0..k}" proof (cases k) case 0 with \f 0 = 0\ show ?thesis by simp next case (Suc k) moreover have "{0..Suc k} = insert 0 {Suc 0..Suc k}" by auto ultimately show ?thesis using \f 0 = 0\ by simp qed end end lemma sum_Suc_diff: fixes f :: "nat \ 'a::ab_group_add" assumes "m \ Suc n" shows "(\i = m..n. f(Suc i) - f i) = f (Suc n) - f m" using assms by (induct n) (auto simp: le_Suc_eq) lemma sum_Suc_diff': fixes f :: "nat \ 'a::ab_group_add" assumes "m \ n" shows "(\i = m..Telescoping\ lemma sum_telescope: fixes f::"nat \ 'a::ab_group_add" shows "sum (\i. f i - f (Suc i)) {.. i} = f 0 - f (Suc i)" by (induct i) simp_all lemma sum_telescope'': assumes "m \ n" shows "(\k\{Suc m..n}. f k - f (k - 1)) = f n - (f m :: 'a :: ab_group_add)" by (rule dec_induct[OF assms]) (simp_all add: algebra_simps) lemma sum_lessThan_telescope: "(\nnThe formula for geometric sums\ lemma sum_power2: "(\i=0.. 1" shows "(\i 0" by simp_all moreover have "(\iy \ 0\) ultimately show ?thesis by simp qed lemma diff_power_eq_sum: fixes y :: "'a::{comm_ring,monoid_mult}" shows "x ^ (Suc n) - y ^ (Suc n) = (x - y) * (\pppp \\COMPLEX_POLYFUN\ in HOL Light\ fixes x :: "'a::{comm_ring,monoid_mult}" shows "x^n - y^n = (x - y) * (\iiiii\n. x^i) = 1 - x^Suc n" by (simp only: one_diff_power_eq lessThan_Suc_atMost) lemma sum_power_shift: fixes x :: "'a::{comm_ring,monoid_mult}" assumes "m \ n" shows "(\i=m..n. x^i) = x^m * (\i\n-m. x^i)" proof - have "(\i=m..n. x^i) = x^m * (\i=m..n. x^(i-m))" by (simp add: sum_distrib_left power_add [symmetric]) also have "(\i=m..n. x^(i-m)) = (\i\n-m. x^i)" using \m \ n\ by (intro sum.reindex_bij_witness[where j="\i. i - m" and i="\i. i + m"]) auto finally show ?thesis . qed lemma sum_gp_multiplied: fixes x :: "'a::{comm_ring,monoid_mult}" assumes "m \ n" shows "(1 - x) * (\i=m..n. x^i) = x^m - x^Suc n" proof - have "(1 - x) * (\i=m..n. x^i) = x^m * (1 - x) * (\i\n-m. x^i)" by (metis mult.assoc mult.commute assms sum_power_shift) also have "... =x^m * (1 - x^Suc(n-m))" by (metis mult.assoc sum_gp_basic) also have "... = x^m - x^Suc n" using assms by (simp add: algebra_simps) (metis le_add_diff_inverse power_add) finally show ?thesis . qed lemma sum_gp: fixes x :: "'a::{comm_ring,division_ring}" shows "(\i=m..n. x^i) = (if n < m then 0 else if x = 1 then of_nat((n + 1) - m) else (x^m - x^Suc n) / (1 - x))" using sum_gp_multiplied [of m n x] apply auto by (metis eq_iff_diff_eq_0 mult.commute nonzero_divide_eq_eq) subsubsection\Geometric progressions\ lemma sum_gp0: fixes x :: "'a::{comm_ring,division_ring}" shows "(\i\n. x^i) = (if x = 1 then of_nat(n + 1) else (1 - x^Suc n) / (1 - x))" using sum_gp_basic[of x n] by (simp add: mult.commute field_split_simps) lemma sum_power_add: fixes x :: "'a::{comm_ring,monoid_mult}" shows "(\i\I. x^(m+i)) = x^m * (\i\I. x^i)" by (simp add: sum_distrib_left power_add) lemma sum_gp_offset: fixes x :: "'a::{comm_ring,division_ring}" shows "(\i=m..m+n. x^i) = (if x = 1 then of_nat n + 1 else x^m * (1 - x^Suc n) / (1 - x))" using sum_gp [of x m "m+n"] by (auto simp: power_add algebra_simps) lemma sum_gp_strict: fixes x :: "'a::{comm_ring,division_ring}" shows "(\iThe formulae for arithmetic sums\ context comm_semiring_1 begin lemma double_gauss_sum: "2 * (\i = 0..n. of_nat i) = of_nat n * (of_nat n + 1)" by (induct n) (simp_all add: sum.atLeast0_atMost_Suc algebra_simps left_add_twice) lemma double_gauss_sum_from_Suc_0: "2 * (\i = Suc 0..n. of_nat i) = of_nat n * (of_nat n + 1)" proof - have "sum of_nat {Suc 0..n} = sum of_nat (insert 0 {Suc 0..n})" by simp also have "\ = sum of_nat {0..n}" by (cases n) (simp_all add: atLeast0_atMost_Suc_eq_insert_0) finally show ?thesis by (simp add: double_gauss_sum) qed lemma double_arith_series: "2 * (\i = 0..n. a + of_nat i * d) = (of_nat n + 1) * (2 * a + of_nat n * d)" proof - have "(\i = 0..n. a + of_nat i * d) = ((\i = 0..n. a) + (\i = 0..n. of_nat i * d))" by (rule sum.distrib) also have "\ = (of_nat (Suc n) * a + d * (\i = 0..n. of_nat i))" by (simp add: sum_distrib_left algebra_simps) finally show ?thesis by (simp add: algebra_simps double_gauss_sum left_add_twice) qed end context unique_euclidean_semiring_with_nat begin lemma gauss_sum: "(\i = 0..n. of_nat i) = of_nat n * (of_nat n + 1) div 2" using double_gauss_sum [of n, symmetric] by simp lemma gauss_sum_from_Suc_0: "(\i = Suc 0..n. of_nat i) = of_nat n * (of_nat n + 1) div 2" using double_gauss_sum_from_Suc_0 [of n, symmetric] by simp lemma arith_series: "(\i = 0..n. a + of_nat i * d) = (of_nat n + 1) * (2 * a + of_nat n * d) div 2" using double_arith_series [of a d n, symmetric] by simp end lemma gauss_sum_nat: "\{0..n} = (n * Suc n) div 2" using gauss_sum [of n, where ?'a = nat] by simp lemma arith_series_nat: "(\i = 0..n. a + i * d) = Suc n * (2 * a + n * d) div 2" using arith_series [of a d n] by simp lemma Sum_Icc_int: "\{m..n} = (n * (n + 1) - m * (m - 1)) div 2" if "m \ n" for m n :: int using that proof (induct i \ "nat (n - m)" arbitrary: m n) case 0 then have "m = n" by arith then show ?case by (simp add: algebra_simps mult_2 [symmetric]) next case (Suc i) have 0: "i = nat((n-1) - m)" "m \ n-1" using Suc(2,3) by arith+ have "\ {m..n} = \ {m..1+(n-1)}" by simp also have "\ = \ {m..n-1} + n" using \m \ n\ by(subst atLeastAtMostPlus1_int_conv) simp_all also have "\ = ((n-1)*(n-1+1) - m*(m-1)) div 2 + n" by(simp add: Suc(1)[OF 0]) also have "\ = ((n-1)*(n-1+1) - m*(m-1) + 2*n) div 2" by simp also have "\ = (n*(n+1) - m*(m-1)) div 2" by (simp add: algebra_simps mult_2_right) finally show ?case . qed lemma Sum_Icc_nat: "\{m..n} = (n * (n + 1) - m * (m - 1)) div 2" for m n :: nat proof (cases "m \ n") case True then have *: "m * (m - 1) \ n * (n + 1)" by (meson diff_le_self order_trans le_add1 mult_le_mono) have "int (\{m..n}) = (\{int m..int n})" by (simp add: sum.atLeast_int_atMost_int_shift) also have "\ = (int n * (int n + 1) - int m * (int m - 1)) div 2" using \m \ n\ by (simp add: Sum_Icc_int) also have "\ = int ((n * (n + 1) - m * (m - 1)) div 2)" using le_square * by (simp add: algebra_simps of_nat_div of_nat_diff) finally show ?thesis by (simp only: of_nat_eq_iff) next case False then show ?thesis by (auto dest: less_imp_Suc_add simp add: not_le algebra_simps) qed lemma Sum_Ico_nat: "\{m..Division remainder\ lemma range_mod: fixes n :: nat assumes "n > 0" shows "range (\m. m mod n) = {0.. ?A \ m \ ?B" proof assume "m \ ?A" with assms show "m \ ?B" by auto next assume "m \ ?B" moreover have "m mod n \ ?A" by (rule rangeI) ultimately show "m \ ?A" by simp qed qed subsection \Products indexed over intervals\ syntax (ASCII) "_from_to_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(PROD _ = _.._./ _)" [0,0,0,10] 10) "_from_upto_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(PROD _ = _..<_./ _)" [0,0,0,10] 10) "_upt_prod" :: "idt \ 'a \ 'b \ 'b" ("(PROD _<_./ _)" [0,0,10] 10) "_upto_prod" :: "idt \ 'a \ 'b \ 'b" ("(PROD _<=_./ _)" [0,0,10] 10) syntax (latex_prod output) "_from_to_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\prod_{\_ = _\<^latex>\}^{\_\<^latex>\}$\ _)" [0,0,0,10] 10) "_from_upto_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\prod_{\_ = _\<^latex>\}^{<\_\<^latex>\}$\ _)" [0,0,0,10] 10) "_upt_prod" :: "idt \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\prod_{\_ < _\<^latex>\}$\ _)" [0,0,10] 10) "_upto_prod" :: "idt \ 'a \ 'b \ 'b" ("(3\<^latex>\$\\prod_{\_ \ _\<^latex>\}$\ _)" [0,0,10] 10) syntax "_from_to_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _.._./ _)" [0,0,0,10] 10) "_from_upto_prod" :: "idt \ 'a \ 'a \ 'b \ 'b" ("(3\_ = _..<_./ _)" [0,0,0,10] 10) "_upt_prod" :: "idt \ 'a \ 'b \ 'b" ("(3\_<_./ _)" [0,0,10] 10) "_upto_prod" :: "idt \ 'a \ 'b \ 'b" ("(3\_\_./ _)" [0,0,10] 10) translations "\x=a..b. t" \ "CONST prod (\x. t) {a..b}" "\x=a.. "CONST prod (\x. t) {a..i\n. t" \ "CONST prod (\i. t) {..n}" "\i "CONST prod (\i. t) {..{int i..int (i+j)}" by (induct j) (auto simp add: atLeastAtMostSuc_conv atLeastAtMostPlus1_int_conv) lemma prod_int_eq: "prod int {i..j} = \{int i..int j}" proof (cases "i \ j") case True then show ?thesis by (metis le_iff_add prod_int_plus_eq) next case False then show ?thesis by auto qed subsection \Efficient folding over intervals\ function fold_atLeastAtMost_nat where [simp del]: "fold_atLeastAtMost_nat f a (b::nat) acc = (if a > b then acc else fold_atLeastAtMost_nat f (a+1) b (f a acc))" by pat_completeness auto termination by (relation "measure (\(_,a,b,_). Suc b - a)") auto lemma fold_atLeastAtMost_nat: assumes "comp_fun_commute f" shows "fold_atLeastAtMost_nat f a b acc = Finite_Set.fold f acc {a..b}" using assms proof (induction f a b acc rule: fold_atLeastAtMost_nat.induct, goal_cases) case (1 f a b acc) interpret comp_fun_commute f by fact show ?case proof (cases "a > b") case True thus ?thesis by (subst fold_atLeastAtMost_nat.simps) auto next case False with 1 show ?thesis by (subst fold_atLeastAtMost_nat.simps) (auto simp: atLeastAtMost_insertL[symmetric] fold_fun_left_comm) qed qed lemma sum_atLeastAtMost_code: "sum f {a..b} = fold_atLeastAtMost_nat (\a acc. f a + acc) a b 0" proof - have "comp_fun_commute (\a. (+) (f a))" by unfold_locales (auto simp: o_def add_ac) thus ?thesis by (simp add: sum.eq_fold fold_atLeastAtMost_nat o_def) qed lemma prod_atLeastAtMost_code: "prod f {a..b} = fold_atLeastAtMost_nat (\a acc. f a * acc) a b 1" proof - have "comp_fun_commute (\a. (*) (f a))" by unfold_locales (auto simp: o_def mult_ac) thus ?thesis by (simp add: prod.eq_fold fold_atLeastAtMost_nat o_def) qed (* TODO: Add support for folding over more kinds of intervals here *) end