diff --git a/IsarMathLib/MetricSpace_ZF.thy b/IsarMathLib/MetricSpace_ZF.thy index d57f8f9..f63795f 100644 --- a/IsarMathLib/MetricSpace_ZF.thy +++ b/IsarMathLib/MetricSpace_ZF.thy @@ -65,7 +65,7 @@ text\A disk is defined as set of points located less than the radius from definition "Disk(X,d,r,c,R) \ {x\X. \d`\c,x\,R\ \ StrictVersion(r)}" -text\We define a metric topology as consisting of unions of open disks.\ +text\We define \metric topology\ as consisting of unions of open disks.\ definition "MetricTopology(X,L,A,r,d) \ {\\. \ \ Pow(\c\X. {Disk(X,d,r,c,R). R\PositiveSet(L,A,r)})}" @@ -94,7 +94,7 @@ locale pmetric_space = loop1 + text\ The next lemma shows the definition of the pseudometric in the notation used in the - \metric_space\ context.\ + \pmetric_space\ context.\ lemma (in pmetric_space) pmetric_properties: shows "d: X\X \ L\<^sup>+" @@ -104,7 +104,7 @@ lemma (in pmetric_space) pmetric_properties: shows using pmetricAssum unfolding IsApseudoMetric_def by auto text\The values of the metric are in the in the nonnegative set of the loop, - hence in the loop. /\ + hence in the loop.\ lemma (in pmetric_space) pmetric_loop_valued: assumes "x\X" "y\X" shows "d`\x,y\ \ L\<^sup>+" "d`\x,y\ \ L" @@ -231,9 +231,9 @@ text\Disks centered at points farther away than the sum of radii do not ov lemma (in pmetric_space) far_disks: assumes "x\X" "y\X" "r\<^sub>x\r\<^sub>y \ d`\x,y\" - shows "disk(x,r\<^sub>x)\disk(y,r\<^sub>y) = 0" + shows "disk(x,r\<^sub>x)\disk(y,r\<^sub>y) = \" proof - - { assume "disk(x,r\<^sub>x)\disk(y,r\<^sub>y) \ 0" + { assume "disk(x,r\<^sub>x)\disk(y,r\<^sub>y) \ \" then obtain z where "z \ disk(x,r\<^sub>x)\disk(y,r\<^sub>y)" by auto then have "z\X" and "d`\x,z\ \ d`\y,z\ \ r\<^sub>x\r\<^sub>y" using disk_definition add_ineq_strict by auto @@ -272,9 +272,9 @@ proof - thus ?thesis by simp qed -text\Unions of disks form a topology, hence (pseudo)metric spaces are topological spaces. - Recall that in the \pmetric_space\ context $\tau$ is the metric topology (i.e. set of unions - of open disks. \ +text\If the order of the loop down-directs its set of positive elements + then the metric topology defined as collection of unions of (open) disks is indeed a topology. + Recall that in the \pmetric_space\ context $\tau$ denotes the metric topology. \ theorem (in pmetric_space) pmetric_is_top: assumes "r {down-directs} L\<^sub>+" @@ -290,7 +290,7 @@ theorem (in pmetric_space) disks_are_base: shows "B {is a base for} \" using assms disks_form_base Top_1_2_T1 metric_top_def_alt by simp -text\$X$ is the carrier of metric topology.\ +text\If $r$ down-directs $L_+$ then $X$ is the carrier of metric topology.\ theorem (in pmetric_space) metric_top_carrier: assumes "r {down-directs} L\<^sub>+" shows "\\ = X" @@ -315,7 +315,7 @@ lemma (in pmetric_space) topology0_valid_in_pmetric_space: shows "topology0(\)" using assms pmetric_is_top unfolding topology0_def by simp -text\Disks are open in the metric topology.\ +text\If $r$ down-directs $L_+$ then disks are open in the metric topology.\ lemma (in pmetric_space) disks_open: assumes "c\X" "R\L\<^sub>+" "r {down-directs} L\<^sub>+" @@ -346,7 +346,7 @@ proof - using ident_indisc posset_definition posset_definition1 by auto qed -text\An ordered loop valued metric space is $T_2$ (i.e. Hausdorff).\ +text\If $r$ down-directs $L_+$ then the ordered loop valued metric space is $T_2$ (i.e. Hausdorff).\ theorem (in metric_space) metric_space_T2: assumes "r {down-directs} L\<^sub>+" @@ -444,9 +444,8 @@ lemma (in pmetric_space) gauge_members: by simp text\Suppose $b\in L^+$ (i.e. b is an element of the loop that is greater than the neutral element) - and $x\in X$. Then the set $B=\{ d^{-1}(\{c\in L^+: c\leq b\}$ is a relation on $X$ and - the image of the singleton set $\{ x\}$ by that relation is the set - $\{ y\in X:d\langle x,y\rangle \leq b\}$, + and $x\in X$. Then the image of the singleton set $\{ x\}$ by the relation + $B=\{ d^{-1}(\{c\in L^+: c\leq b\}$ is the set $\{ y\in X:d\langle x,y\rangle \leq b\}$, i.e. the closed disk with center $x$ and radius $b$. Hence the the image $B\{ x\}$ contains the open disk with center $x$ and radius $b$. \ @@ -518,7 +517,8 @@ corollary (in pmetric_space) gauge_3rd_cond: assumes "B\<^sub>1\\" shows "\B\<^sub>2\\. B\<^sub>2 \ converse(B\<^sub>1)" using assms gauge_symmetric by auto -text\The sets of the form $d^{-1}([0,b])$ are subsets of $X\times X$. \ +text\The collection of sets of the form $d^{-1}([0,b])$ for $b\in L_+$ + is contained of the powerset of $X\times X$.\ lemma (in pmetric_space) gauge_5thCond: shows "\\Pow(X\X)" using uniform_gauge_def_alt pmetric_properties(1) func1_1_L3 by force diff --git a/IsarMathLib/document/root.tex b/IsarMathLib/document/root.tex index a20e070..c7029ed 100644 --- a/IsarMathLib/document/root.tex +++ b/IsarMathLib/document/root.tex @@ -66,7 +66,7 @@ \maketitle \begin{abstract} -This is the proof document of the IsarMathLib project version 1.30.0. +This is the proof document of the IsarMathLib project version 1.31.0. IsarMathLib is a library of formalized mathematics for Isabelle2024 (ZF logic). \end{abstract} diff --git a/docs/IsarMathLib/Fol1.html b/docs/IsarMathLib/Fol1.html index 7a27d20..da3afb8 100644 --- a/docs/IsarMathLib/Fol1.html +++ b/docs/IsarMathLib/Fol1.html @@ -193,20 +193,20 @@

Theory Fol1

using
Xor_def by auto text‹Constructions from the same sets are the same. - It is suprising but we do have to use this as a rule in rarte cases.› + It is suprising but we do have to use this as a rule in rare cases.› -lemma same_constr: assumes "x=y" shows "P(x) = P(y)" - using assms by simp +lemma same_constr: assumes "x=y" shows "P(x) = P(y)" + using assms by simp text‹Equivalence relations are symmetric.› -lemma equiv_is_sym: assumes A1: "equiv(X,r)" and A2: "x,y r" +lemma equiv_is_sym: assumes A1: "equiv(X,r)" and A2: "x,y r" shows "y,x r" proof - - from A1 have "sym(r)" using equiv_def by simp + from A1 have "sym(r)" using equiv_def by simp then have "x y. x,y r y,x r" unfolding sym_def by fast - with A2 show "y,x r" by blast + with A2 show "y,x r" by blast qed (* In Isabelle/ZF conjunction associates to the right!. diff --git a/docs/IsarMathLib/Group_ZF_5.html b/docs/IsarMathLib/Group_ZF_5.html index 6753ac6..b75b529 100644 --- a/docs/IsarMathLib/Group_ZF_5.html +++ b/docs/IsarMathLib/Group_ZF_5.html @@ -63,53 +63,55 @@

Theory Group_ZF_5

"IsMorphism(G,P,F,f) g1G. g2G. f`(P`g1,g2) = F`f`(g1),f`(g2)" text‹A function $f:G\rightarrow H$ between algebraic structures - $(G,\cdot_G)$ and $(H,\cdot_H)$ with one operation (each) is a homomorphism is it has the morphism - property. › + $(G,\cdot_G)$ and $(H,\cdot_H)$ with one operation (each) is a homomorphism if + it has the morphism property. ›
definition "Homomor(f,G,P,H,F) f:GH IsMorphism(G,P,F,f)" text‹Now a lemma about the definition:› -lemma homomor_eq: +lemma homomor_eq: assumes "Homomor(f,G,P,H,F)" "g1G" "g2G" shows "f`(P`g1,g2) = F`f`(g1),f`(g2)" - using assms unfolding Homomor_def IsMorphism_def by auto + using assms unfolding Homomor_def IsMorphism_def by auto -text‹An endomorphism is a homomorphism from a group to the same group. In case -the group is abelian, it has a nice structure.› +text‹An endomorphism is a homomorphism from a group to the same group. + We define End(G,P)› as the set of endomorphisms for a given group. + As we show later when the group is abelian, the set of endomorphisms + with pointwise adddition and composition as multiplication forms a ring.› definition "End(G,P) {fGG. Homomor(f,G,P,G,P)}" text‹The defining property of an endomorphism written in notation used in group0› context:› -lemma (in group0) endomor_eq: assumes "f End(G,P)" "g1G" "g2G" +lemma (in group0) endomor_eq: assumes "f End(G,P)" "g1G" "g2G" shows "f`(g1g2) = f`(g1)f`(g2)" - using assms homomor_eq unfolding End_def by auto + using assms homomor_eq unfolding End_def by auto text‹A function that maps a group $G$ into itself and satisfies $f(g_1\cdot g2) = f(g_1)\cdot f(g_2)$ is an endomorphism.› -lemma (in group0) eq_endomor: +lemma (in group0) eq_endomor: assumes "f:GG" and "g1G. g2G. f`(g1g2)=f`(g1)f`(g2)" shows "f End(G,P)" - using assms unfolding End_def Homomor_def IsMorphism_def by simp + using assms unfolding End_def Homomor_def IsMorphism_def by simp text‹The set of endomorphisms forms a submonoid of the monoid of function from a set to that set under composition.› -lemma (in group0) end_composition: +lemma (in group0) end_composition: assumes "f1End(G,P)" "f2End(G,P)" shows "Composition(G)`f1,f2 End(G,P)" proof- - from assms have fun: "f1:GG" "f2:GG" unfolding End_def by auto + from assms have fun: "f1:GG" "f2:GG" unfolding End_def by auto then have "f1 O f2:GG" using comp_fun by auto - from assms fun(2) have + from assms fun(2) have "g1G. g2G. (f1 O f2)`(g1g2) = ((f1 O f2)`(g1))((f1 O f2)`(g2))" - using group_op_closed comp_fun_apply endomor_eq apply_type + using group_op_closed comp_fun_apply endomor_eq apply_type by simp - with fun f1 O f2:GG show ?thesis using eq_endomor func_ZF_5_L2 + with fun f1 O f2:GG show ?thesis using eq_endomor func_ZF_5_L2 by simp qed @@ -122,62 +124,63 @@

Theory Group_ZF_5

where "InEnd(F,G,P) restrict(F,End(G,P)×End(G,P))" text‹Endomoprhisms of a group form a monoid with composition as the binary operation, - with the identity map as the neutral element.› + and the identity map as the neutral element.› -theorem (in group0) end_comp_monoid: +theorem (in group0) end_comp_monoid: shows "IsAmonoid(End(G,P),InEnd(Composition(G),G,P))" and "TheNeutralElement(End(G,P),InEnd(Composition(G),G,P)) = id(G)" proof - let ?C0 = "InEnd(Composition(G),G,P)" - have fun: "id(G):GG" unfolding id_def by auto + have fun: "id(G):GG" unfolding id_def by auto { fix g h assume "gG""hG" then have "id(G)`(gh)=(id(G)`g)(id(G)`h)" using group_op_closed by simp } - with groupAssum fun have "id(G) End(G,P)" using eq_endomor by simp - moreover have A0: "id(G)=TheNeutralElement(G G, Composition(G))" + with groupAssum fun have "id(G) End(G,P)" using eq_endomor by simp + moreover have A0: "id(G)=TheNeutralElement(G G, Composition(G))" using Group_ZF_2_5_L2(2) by auto - ultimately have A1: "TheNeutralElement(G G, Composition(G)) End(G,P)" by auto - moreover have A2: "End(G,P) GG" unfolding End_def by blast - moreover have A3: "End(G,P) {is closed under} Composition(G)" - using end_composition unfolding IsOpClosed_def by blast + ultimately have A1: "TheNeutralElement(G G, Composition(G)) End(G,P)" by auto + moreover have A2: "End(G,P) GG" unfolding End_def by blast + moreover have A3: "End(G,P) {is closed under} Composition(G)" + using end_composition unfolding IsOpClosed_def by blast ultimately show "IsAmonoid(End(G,P),?C0)" using monoid0.group0_1_T1 Group_ZF_2_5_L2(1) unfolding monoid0_def by blast have "IsAmonoid(GG,Composition(G))" using Group_ZF_2_5_L2(1) by auto - with A0 A1 A2 A3 show "TheNeutralElement(End(G,P),?C0) = id(G)" + with A0 A1 A2 A3 show "TheNeutralElement(End(G,P),?C0) = id(G)" using group0_1_L6 by auto qed text‹The set of endomorphisms is closed under pointwise addition (derived from the group operation). This is so because the group is abelian.› -theorem (in abelian_group) end_pointwise_addition: - assumes "fEnd(G,P)" "gEnd(G,P)" "F = P {lifted to function space over} G" +theorem (in abelian_group) end_pointwise_addition: + assumes "fEnd(G,P)" "gEnd(G,P)" + defines "F P {lifted to function space over} G" shows "F`f,g End(G,P)" proof- - from assms(1,2) have fun: "f:GG" "gGG" unfolding End_def by simp_all - with assms(3) have fun2: "F`f,g:GG" + from assms(1,2) have fun: "f:GG" "gGG" unfolding End_def by simp_all + with assms(3) have fun2: "F`f,g:GG" using monoid0.Group_ZF_2_1_L0 group0_2_L1 by simp { fix g1 g2 assume "g1G" "g2G" - with isAbelian assms fun have + with isAbelian assms fun have "(F`f,g)`(g1g2) = (F`f,g)`(g1)(F`f,g)`(g2)" - using Group_ZF_2_1_L3 group_op_closed endomor_eq + using Group_ZF_2_1_L3 group_op_closed endomor_eq apply_type group0_4_L8(3) Group_ZF_2_1_L3 by simp - } with fun2 show ?thesis using eq_endomor by simp + } with fun2 show ?thesis using eq_endomor by simp qed text‹The value of a product of endomorphisms on a group element is the product of values.› -lemma (in abelian_group) end_pointwise_add_val: +lemma (in abelian_group) end_pointwise_add_val: assumes "fEnd(G,P)" "gEnd(G,P)" "xG" "F = P {lifted to function space over} G" shows "(InEnd(F,G,P)`f,g)`(x) = (f`(x))(g`(x))" - using assms group_oper_fun monoid.group0_1_L3B func_ZF_1_L4 + using assms group_oper_fun monoid.group0_1_L3B func_ZF_1_L4 unfolding End_def by simp -text‹The inverse of an abelian group is an endomorphism.› +text‹The operation of taking the inverse in an abelian group is an endomorphism.› -lemma (in abelian_group) end_inverse_group: +lemma (in abelian_group) end_inverse_group: shows "GroupInv(G,P) End(G,P)" using inverse_in_group group_inv_of_two isAbelian IsCommutative_def group0_2_T2 groupAssum Homomor_def @@ -186,47 +189,47 @@

Theory Group_ZF_5

text‹The set of homomorphisms of an abelian group is an abelian subgroup of the group of functions from a set to a group, under pointwise addition.› -theorem (in abelian_group) end_addition_group: +theorem (in abelian_group) end_addition_group: assumes "F = P {lifted to function space over} G" shows "IsAgroup(End(G,P),InEnd(F,G,P))" and "InEnd(F,G,P) {is commutative on} End(G,P)" proof- - have "End(G,P)0" using end_comp_monoid(1) monoid0.group0_1_L3A + have "End(G,P)0" using end_comp_monoid(1) monoid0.group0_1_L3A unfolding monoid0_def by auto moreover have "End(G,P) GG" unfolding End_def by auto - moreover from isAbelian assms(1) have "End(G,P){is closed under} F" - unfolding IsOpClosed_def using end_pointwise_addition by auto - moreover from groupAssum assms(1) have + moreover from isAbelian assms(1) have "End(G,P){is closed under} F" + unfolding IsOpClosed_def using end_pointwise_addition by auto + moreover from groupAssum assms(1) have "fEnd(G,P). GroupInv(GG,F)`(f) End(G,P)" - using monoid0.group0_1_L1 end_composition(1) end_inverse_group + using monoid0.group0_1_L1 end_composition(1) end_inverse_group func_ZF_5_L2 group0_2_T2 Group_ZF_2_1_L6 unfolding monoid0_def End_def by force ultimately show "IsAgroup(End(G,P),InEnd(F,G,P))" - using assms(1) group0.group0_3_T3 Group_ZF_2_1_T2 + using assms(1) group0.group0_3_T3 Group_ZF_2_1_T2 unfolding IsAsubgroup_def group0_def by blast - from assms(1) isAbelian show + from assms(1) isAbelian show "InEnd(F,G,P) {is commutative on} End(G,P)" using Group_ZF_2_1_L7 unfolding End_def IsCommutative_def by auto qed text‹Endomorphisms form a subgroup of the space of functions that map the group to itself.› -lemma (in abelian_group) end_addition_subgroup: +lemma (in abelian_group) end_addition_subgroup: shows "IsAsubgroup(End(G,P),P {lifted to function space over} G)" - using end_addition_group unfolding IsAsubgroup_def by simp + using end_addition_group unfolding IsAsubgroup_def by simp text‹The neutral element of the group of endomorphisms of a group is the constant function with value equal to the neutral element of the group.› -lemma (in abelian_group) end_add_neut_elem: +lemma (in abelian_group) end_add_neut_elem: assumes "F = P {lifted to function space over} G" shows "TheNeutralElement(End(G,P),InEnd(F,G,P)) = ConstantFunction(G,𝟭)" - using assms end_addition_subgroup lift_group_subgr_neut by simp + using assms end_addition_subgroup lift_group_subgr_neut by simp text‹For the endomorphisms of a group $G$ the group operation lifted to the function space over $G$ is distributive with respect to the composition operation. › -lemma (in abelian_group) distributive_comp_pointwise: +lemma (in abelian_group) distributive_comp_pointwise: assumes "F = P {lifted to function space over} G" shows "IsDistributive(End(G,P),InEnd(F,G,P),InEnd(Composition(G),G,P))" @@ -234,70 +237,70 @@

Theory Group_ZF_5

let ?CG = "Composition(G)" let ?CE = "InEnd(?CG,G,P)" let ?FE = "InEnd(F,G,P)" - { fix b c d assume AS: "bEnd(G,P)" "cEnd(G,P)" "dEnd(G,P)" - with assms(1) have ig1: "?CG `b, F ` c, d = b O (F`c,d)" + { fix b c d assume AS: "bEnd(G,P)" "cEnd(G,P)" "dEnd(G,P)" + with assms(1) have ig1: "?CG `b, F ` c, d = b O (F`c,d)" using monoid.Group_ZF_2_1_L0 func_ZF_5_L2 unfolding End_def by auto - with AS have ig2: "F`?CG`b,c,?CG `b,d = F`b O c,b O d" + with AS have ig2: "F`?CG`b,c,?CG `b,d = F`b O c,b O d" unfolding End_def using func_ZF_5_L2 by auto - from assms(1) AS have comp1fun: "(b O (F`c,d)):GG" + from assms(1) AS have comp1fun: "(b O (F`c,d)):GG" using monoid.Group_ZF_2_1_L0 comp_fun unfolding End_def by force - from assms(1) AS have comp2fun: "(F `b O c,b O d) : GG" + from assms(1) AS have comp2fun: "(F `b O c,b O d) : GG" using monoid.Group_ZF_2_1_L0 comp_fun unfolding End_def by force { fix g assume "gG" - with assms(1) AS(2,3) have "(b O (F`c,d))`(g) = b`((F`c,d)`(g))" + with assms(1) AS(2,3) have "(b O (F`c,d))`(g) = b`((F`c,d)`(g))" using comp_fun_apply monoid.Group_ZF_2_1_L0 unfolding End_def by force - with groupAssum assms(1) AS gG have + with groupAssum assms(1) AS gG have "(b O (F`c,d))`(g) = (F`b O c,b O d)`(g)" using Group_ZF_2_1_L3 apply_type homomor_eq comp_fun unfolding End_def by auto } hence "gG. (b O (F`c,d))`(g) = (F`b O c,b O d)`(g)" by simp - with comp1fun comp2fun ig1 ig2 have + with comp1fun comp2fun ig1 ig2 have "?CG`b,F`c, d = F`?CG`b , c,?CG`b,d" using func_eq by simp - moreover from AS(2,3) have "F`c, d = ?FE`c, d" + moreover from AS(2,3) have "F`c, d = ?FE`c, d" using restrict by simp - moreover from AS have "?CG`b,c = ?CE`b,c" and "?CG`b,d = ?CE`b,d" + moreover from AS have "?CG`b,c = ?CE`b,c" and "?CG`b,d = ?CE`b,d" using restrict by auto - moreover from assms AS have "?CG`b,F `c,d = ?CE`b, F`c, d" - using end_pointwise_addition by simp - moreover from AS have "F`?CG`b,c,?CG`b,d = ?FE`?CG `b,c,?CG `b,d" - using end_composition by simp - ultimately have eq1: "?CE`b, ?FE`c,d = ?FE `?CE`b,c,?CE`b,d" + moreover from assms AS have "?CG`b,F `c,d = ?CE`b, F`c, d" + using end_pointwise_addition by simp + moreover from AS have "F`?CG`b,c,?CG`b,d = ?FE`?CG `b,c,?CG `b,d" + using end_composition by simp + ultimately have eq1: "?CE`b, ?FE`c,d = ?FE `?CE`b,c,?CE`b,d" by simp - from assms(1) AS have - compfun: "(F`c,d) O b : GG" "F`c O b,d O b : GG" + from assms(1) AS have + compfun: "(F`c,d) O b : GG" "F`c O b,d O b : GG" using monoid.Group_ZF_2_1_L0 comp_fun unfolding End_def by auto { fix g assume "gG" - with AS(1) have bg: "b`(g) G" unfolding End_def using apply_type + with AS(1) have bg: "b`(g) G" unfolding End_def using apply_type by auto - from gG AS(1) have "((F`c,d) O b)`g = (F`c,d)`(b`(g))" + from gG AS(1) have "((F`c,d) O b)`g = (F`c,d)`(b`(g))" using comp_fun_apply unfolding End_def by force - also from assms(1) AS(2,3) bg have " = (c`(b`(g)))(d`(b`(g)))" + also from assms(1) AS(2,3) bg have " = (c`(b`(g)))(d`(b`(g)))" using Group_ZF_2_1_L3 unfolding End_def by auto - also from gG AS have " = ((c O b)`(g))((d O b)`(g))" + also from gG AS have " = ((c O b)`(g))((d O b)`(g))" using comp_fun_apply unfolding End_def by auto - also from assms(1) gG AS have " = (F`c O b,d O b)`g" + also from assms(1) gG AS have " = (F`c O b,d O b)`g" using comp_fun Group_ZF_2_1_L3 unfolding End_def by auto finally have "((F`c,d) O b)`(g) = (F`c O b,d O b)`(g)" by simp } hence "gG. ((F`c,d) O b)`(g) = (F`c O b,d O b)`(g)" by simp - with compfun have "(F`c,d) O b = F`c O b,d O b" + with compfun have "(F`c,d) O b = F`c O b,d O b" using func_eq by blast - with assms(1) AS have "?CG`F`c,d,b = F`?CG`c,b,?CG`d , b" + with assms(1) AS have "?CG`F`c,d,b = F`?CG`c,b,?CG`d , b" using monoid.Group_ZF_2_1_L0 func_ZF_5_L2 unfolding End_def by simp - moreover from AS(2,3) have "F`c, d = ?FE`c, d" + moreover from AS(2,3) have "F`c, d = ?FE`c, d" using restrict by simp - moreover from AS have "?CG`c,b = ?CE`c , b" "?CG`d,b = ?CE`d,b" + moreover from AS have "?CG`c,b = ?CE`c , b" "?CG`d,b = ?CE`d,b" using restrict by auto - moreover from assms AS have "?CG`F`c,d,b = ?CE`F`c,d,b" - using end_pointwise_addition by auto - moreover from AS have "F`?CG`c,b,?CG`d,b = ?FE`?CG`c,b,?CG`d,b" - using end_composition by auto + moreover from assms AS have "?CG`F`c,d,b = ?CE`F`c,d,b" + using end_pointwise_addition by auto + moreover from AS have "F`?CG`c,b,?CG`d,b = ?FE`?CG`c,b,?CG`d,b" + using end_composition by auto ultimately have "?CE`?FE`c,d,b = ?FE`?CE`c,b,?CE`d,b" by auto - with eq1 have "(?CE`b, ?FE`c, d = ?FE`?CE`b,c,?CE`b,d) + with eq1 have "(?CE`b, ?FE`c, d = ?FE`?CE`b,c,?CE`b,d) (?CE`?FE`c,d,b = ?FE`?CE`c,b,?CE`d,b)" by auto } then show ?thesis unfolding IsDistributive_def by auto @@ -306,11 +309,11 @@

Theory Group_ZF_5

text‹The endomorphisms of an abelian group is in fact a ring with the previous operations.› -theorem (in abelian_group) end_is_ring: +theorem (in abelian_group) end_is_ring: assumes "F = P {lifted to function space over} G" shows "IsAring(End(G,P),InEnd(F,G,P),InEnd(Composition(G),G,P))" - using assms end_addition_group end_comp_monoid(1) distributive_comp_pointwise + using assms end_addition_group end_comp_monoid(1) distributive_comp_pointwise unfolding IsAring_def by auto text‹The theorems proven in the ring0› context are valid in the abelian_group› context @@ -330,7 +333,7 @@

Theory Group_ZF_5

TheNeutralElement (End(G, P), InEnd(Composition(G),G,P)), TheNeutralElement (End(G, P), InEnd(Composition(G),G,P))"
"λx. InEnd(Composition(G),G,P)`x, x" - using end_is_ring unfolding ring0_def by blast + using end_is_ring unfolding ring0_def by blast subsection‹First isomorphism theorem› @@ -339,80 +342,80 @@

Theory Group_ZF_5

text‹A group homomorphism sends the neutral element to the neutral element.› -lemma image_neutral: +lemma image_neutral: assumes "IsAgroup(G,P)" "IsAgroup(H,F)" "Homomor(f,G,P,H,F)" shows "f`(TheNeutralElement(G,P)) = TheNeutralElement(H,F)" proof - let ?eG = "TheNeutralElement(G,P)" let ?eH = "TheNeutralElement(H,F)" - from assms(3) have ff: "f:GH" + from assms(3) have ff: "f:GH" unfolding Homomor_def by simp - have g: "?eG = P`?eG,?eG" "?eG G" - using assms(1) group0.group0_2_L2 unfolding group0_def by simp_all - with assms have "f`(?eG) = F`f`(?eG),f`(?eG)" + have g: "?eG = P`?eG,?eG" "?eG G" + using assms(1) group0.group0_2_L2 unfolding group0_def by simp_all + with assms have "f`(?eG) = F`f`(?eG),f`(?eG)" unfolding Homomor_def IsMorphism_def by force moreover - from ff g(2) have h: "f`(?eG) H" using apply_type + from ff g(2) have h: "f`(?eG) H" using apply_type by simp - with assms(2) have "f`(?eG) = F`f`(?eG),?eH" + with assms(2) have "f`(?eG) = F`f`(?eG),?eH" using group0.group0_2_L2 unfolding group0_def by simp ultimately have "F`f`(?eG),?eH = F`f`(?eG),f`(?eG)" by simp - with assms(2) h have + with assms(2) h have "LeftTranslation(H,F,f`(?eG))`(?eH) = LeftTranslation(H,F,f`(?eG))`(f`(?eG))" using group0.group0_5_L2(2) group0.group0_2_L2 unfolding group0_def by simp - moreover from assms(2) h have "LeftTranslation(H,F,f`(?eG))inj(H,H)" + moreover from assms(2) h have "LeftTranslation(H,F,f`(?eG))inj(H,H)" using group0.trans_bij(2) unfolding group0_def bij_def by simp - ultimately show ?thesis using h assms(2) group0.group0_2_L2 + ultimately show ?thesis using h assms(2) group0.group0_2_L2 unfolding inj_def group0_def by force qed text‹If $f:G\rightarrow H$ is a homomorphism, then it commutes with the inverse › -lemma image_inv: +lemma image_inv: assumes "IsAgroup(G,P)" "IsAgroup(H,F)" "Homomor(f,G,P,H,F)" "gG" shows "f`(GroupInv(G,P)`(g)) = GroupInv(H,F)`(f`(g))" proof - - from assms(3) have ff: "f:GH" + from assms(3) have ff: "f:GH" unfolding Homomor_def by simp - with assms(4) have im: "f`(g)H" using apply_type by simp - from assms(1,4) have inv: "GroupInv(G,P)`(g)G" + with assms(4) have im: "f`(g)H" using apply_type by simp + from assms(1,4) have inv: "GroupInv(G,P)`(g)G" using group0.inverse_in_group unfolding group0_def by simp - with ff have inv2: "f`(GroupInv(G,P)`g)H" using apply_type by simp - from assms(1,4) have + with ff have inv2: "f`(GroupInv(G,P)`g)H" using apply_type by simp + from assms(1,4) have "f`(TheNeutralElement(G,P)) = f`(P`g,GroupInv(G,P)`(g))" using group0.group0_2_L6 unfolding group0_def by simp - also from assms inv have " = F`f`(g),f`(GroupInv(G,P)`(g))" + also from assms inv have " = F`f`(g),f`(GroupInv(G,P)`(g))" unfolding Homomor_def IsMorphism_def by simp finally have "f`(TheNeutralElement(G,P)) = F`f`(g),f`(GroupInv(G,P)`(g))" by simp - with assms im inv2 show ?thesis + with assms im inv2 show ?thesis using group0.group0_2_L9 image_neutral unfolding group0_def by simp qed text‹The preimage of a subgroup is a subgroup› -theorem preimage_sub: +theorem preimage_sub: assumes "IsAgroup(G,P)" "IsAgroup(H,F)" "Homomor(f,G,P,H,F)" "IsAsubgroup(K,F)" shows "IsAsubgroup(f-``(K),P)" proof - - from assms(3) have ff: "f:GH" + from assms(3) have ff: "f:GH" unfolding Homomor_def by simp - from assms(2) have Hgr: "group0(H,F)" unfolding group0_def by simp - from assms(1) have Ggr: "group0(G,P)" unfolding group0_def by simp + from assms(2) have Hgr: "group0(H,F)" unfolding group0_def by simp + from assms(1) have Ggr: "group0(G,P)" unfolding group0_def by simp moreover - from assms ff Ggr Hgr have "TheNeutralElement(G,P) f-``(K)" + from assms ff Ggr Hgr have "TheNeutralElement(G,P) f-``(K)" using image_neutral group0.group0_3_L5 func1_1_L15 group0.group0_2_L2 by simp hence "f-``(K)0" by blast - moreover from ff have "f-``(K) G" using func1_1_L3 by simp - moreover from assms ff Ggr Hgr have "f-``(K) {is closed under} P" + moreover from ff have "f-``(K) G" using func1_1_L3 by simp + moreover from assms ff Ggr Hgr have "f-``(K) {is closed under} P" using func1_1_L15 group0.group0_3_L6 group0.group_op_closed func1_1_L15 unfolding IsOpClosed_def Homomor_def IsMorphism_def by simp - moreover from assms ff Ggr Hgr have + moreover from assms ff Ggr Hgr have "xf-``(K). GroupInv(G, P)`(x) f-``(K)" using group0.group0_3_T3A image_inv func1_1_L15 group0.inverse_in_group by simp @@ -421,42 +424,42 @@

Theory Group_ZF_5

text‹The preimage of a normal subgroup is normal› -theorem preimage_normal_subgroup: +theorem preimage_normal_subgroup: assumes "IsAgroup(G,P)" "IsAgroup(H,F)" "Homomor(f,G,P,H,F)" "IsAnormalSubgroup(H,F,K)" shows "IsAnormalSubgroup(G,P,f-``(K))" proof - - from assms(3) have ff: "f:GH" + from assms(3) have ff: "f:GH" unfolding Homomor_def by simp - from assms(2) have Hgr: "group0(H,F)" unfolding group0_def by simp - with assms(4) have "KH" using group0.group0_3_L2 + from assms(2) have Hgr: "group0(H,F)" unfolding group0_def by simp + with assms(4) have "KH" using group0.group0_3_L2 unfolding IsAnormalSubgroup_def by simp - from assms(1) have Ggr: "group0(G,P)" unfolding group0_def by simp - moreover from assms have "IsAsubgroup(f-``(K),P)" + from assms(1) have Ggr: "group0(G,P)" unfolding group0_def by simp + moreover from assms have "IsAsubgroup(f-``(K),P)" using preimage_sub unfolding IsAnormalSubgroup_def by simp moreover - { fix g assume gG: "gG" + { fix g assume gG: "gG" { fix h assume "h {P`g,P`h, GroupInv(G, P)`(g). h f-``(K)}" then obtain k where - k: "h = P`g,P`k,GroupInv(G, P)`(g)" "k f-``(K)" + k: "h = P`g,P`k,GroupInv(G, P)`(g)" "k f-``(K)" by auto - from k(1) have "f`(h) = f`(P`g,P`k, GroupInv(G, P)`(g))" by simp - moreover from ff k(2) have "kG" using vimage_iff + from k(1) have "f`(h) = f`(P`g,P`k, GroupInv(G, P)`(g))" by simp + moreover from ff k(2) have "kG" using vimage_iff unfolding Pi_def by blast - ultimately have f: "f`(h) = F`f`(g),F`f`(k),GroupInv(H,F)`(f`(g))" - using assms(1-4) Ggr gG group0.group_op_closed + ultimately have f: "f`(h) = F`f`(g),F`f`(k),GroupInv(H,F)`(f`(g))" + using assms(1-4) Ggr gG group0.group_op_closed group0.inverse_in_group image_inv homomor_eq by simp - from assms(1) ff Ggr gG k have "hG" using group0.group_op_closed + from assms(1) ff Ggr gG k have "hG" using group0.group_op_closed group0.inverse_in_group func1_1_L15 by simp - from assms(4) ff k(2) gG have "f`(k)K" "f`(g)H" and + from assms(4) ff k(2) gG have "f`(k)K" "f`(g)H" and "F`F`f`(g),f`(k),GroupInv(H,F)`(f`(g)) K" using func1_1_L15 apply_type unfolding IsAnormalSubgroup_def by auto - moreover from f`(k)K KH Hgr f f`(g)H have + moreover from f`(k)K KH Hgr f f`(g)H have "f`(h) = F`F`f`(g),f`(k),GroupInv(H,F)`(f`(g))" using group0.group_oper_assoc group0.inverse_in_group by auto ultimately have "f`(h) K" by simp - with ff hG have "h f-``(K)" using func1_1_L15 by simp + with ff hG have "h f-``(K)" using func1_1_L15 by simp } hence "{P`g,P`h,GroupInv(G,P)`(g). hf-``(K)} f-``(K)" by blast } hence "gG. {P`g, P`h, GroupInv(G, P)`(g). hf-``(K)} f-``(K)" @@ -466,50 +469,50 @@

Theory Group_ZF_5

text‹The kernel of an homomorphism is a normal subgroup.› -corollary kernel_normal_sub: +corollary kernel_normal_sub: assumes "IsAgroup(G,P)" "IsAgroup(H,F)" "Homomor(f,G,P,H,F)" shows "IsAnormalSubgroup(G,P,f-``{TheNeutralElement(H,F)})" - using assms preimage_normal_subgroup group0.trivial_normal_subgroup + using assms preimage_normal_subgroup group0.trivial_normal_subgroup unfolding group0_def by auto text‹The image of a subgroup is a subgroup› -theorem image_subgroup: +theorem image_subgroup: assumes "IsAgroup(G,P)" "IsAgroup(H,F)" "Homomor(f,G,P,H,F)" "f:GH" "IsAsubgroup(K,P)" shows "IsAsubgroup(f``K,F)" proof - - from assms(1,5) have sub: "KG" using group0.group0_3_L2 + from assms(1,5) have sub: "KG" using group0.group0_3_L2 unfolding group0_def by simp - from assms(2) have "group0(H,F)" unfolding group0_def by simp - moreover from assms(4) have "f``(K) H" - using func_imagedef sub apply_type by auto + from assms(2) have "group0(H,F)" unfolding group0_def by simp + moreover from assms(4) have "f``(K) H" + using func_imagedef sub apply_type by auto moreover - from assms(1,4,5) sub have "f`(TheNeutralElement(G,P)) f``(K)" + from assms(1,4,5) sub have "f`(TheNeutralElement(G,P)) f``(K)" using group0.group0_3_L5 func_imagedef unfolding group0_def by auto hence "f``(K) 0" by blast moreover { fix x assume "xf``(K)" - with assms(4) sub obtain q where q: "qK" "x=f`(q)" + with assms(4) sub obtain q where q: "qK" "x=f`(q)" using func_imagedef by auto - with assms(1-4) sub have "GroupInv(H,F)`(x) = f`(GroupInv(G,P)`q)" + with assms(1-4) sub have "GroupInv(H,F)`(x) = f`(GroupInv(G,P)`q)" using image_inv by auto - with assms(1,4,5) q(1) sub have "GroupInv(H,F)`(x) f``(K)" + with assms(1,4,5) q(1) sub have "GroupInv(H,F)`(x) f``(K)" using group0.group0_3_T3A func_imagedef unfolding group0_def by auto } hence "xf``(K). GroupInv(H, F)`(x) f``(K)" by auto moreover { fix x y assume "xf``(K)" "yf``(K)" - with assms(4) sub obtain qx qy where - q: "qxK" "x=f`(qx)" "qyK" "y=f`(qy)" + with assms(4) sub obtain qx qy where + q: "qxK" "x=f`(qx)" "qyK" "y=f`(qy)" using func_imagedef by auto - with assms(1-3) sub have "F`x,y = f`(P`qx,qy)" + with assms(1-3) sub have "F`x,y = f`(P`qx,qy)" using homomor_eq by force - moreover from assms(1,5) q(1,3) have "P`qx,qy K" + moreover from assms(1,5) q(1,3) have "P`qx,qy K" using group0.group0_3_L6 unfolding group0_def by simp ultimately have "F`x,y f``(K)" - using assms(4) sub func_imagedef by auto + using assms(4) sub func_imagedef by auto } then have "f``(K) {is closed under} F" unfolding IsOpClosed_def by simp ultimately show ?thesis using group0.group0_3_T3 by simp @@ -517,14 +520,14 @@

Theory Group_ZF_5

text‹The image of a group under a homomorphism is a subgroup of the target group.› -corollary image_group: +corollary image_group: assumes "IsAgroup(G,P)" "IsAgroup(H,F)" "Homomor(f,G,P,H,F)" shows "IsAsubgroup(f``(G),F)" proof - - from assms(1) have "restrict(P,G×G) = P" + from assms(1) have "restrict(P,G×G) = P" using group0.group_oper_fun restrict_domain unfolding group0_def by blast - with assms show ?thesis using image_subgroup + with assms show ?thesis using image_subgroup unfolding Homomor_def IsAsubgroup_def by simp qed @@ -532,132 +535,132 @@

Theory Group_ZF_5

that any group homomorphism $f:G\to H$ gives an isomorphism between a quotient group of $G$ and a subgroup of $H$.›
-theorem isomorphism_first_theorem: +theorem isomorphism_first_theorem: assumes "IsAgroup(G,P)" "IsAgroup(H,F)" "Homomor(f,G,P,H,F)" defines "r QuotientGroupRel(G,P,f-``{TheNeutralElement(H,F)})" and "𝒫 QuotientGroupOp(G,P,f-``{TheNeutralElement(H,F)})" shows "𝔣. Homomor(𝔣,G//r,𝒫,f``(G),restrict(F,(f``(G))×(f``(G)))) 𝔣bij(G//r,f``(G))" proof- let ?𝔣 = "{r``{g},f`(g). gG}" - from assms(3) have ff: "f:GH" + from assms(3) have ff: "f:GH" unfolding Homomor_def by simp - from assms(1-5) have "equiv(G,r)" + from assms(1-5) have "equiv(G,r)" using group0.Group_ZF_2_4_L3 kernel_normal_sub unfolding group0_def IsAnormalSubgroup_def by simp - from assms(4) ff have "?𝔣 Pow((G//r)×f``(G))" + from assms(4) ff have "?𝔣 Pow((G//r)×f``(G))" unfolding quotient_def using func_imagedef by auto moreover have "(G//r) domain(?𝔣)" unfolding domain_def quotient_def by auto moreover { fix x y t assume A: "x,y ?𝔣" "x,t ?𝔣" then obtain gy gr where "x, y=r``{gy},f`(gy)" "x, t=r``{gr},f`(gr)" and "grG" "gyG" by auto - hence B: "r``{gy}=r``{gr}" "y=f`(gy)" "t=f`(gr)" by auto - from ff gyG grG B(2,3) have "yH" "tH" + hence B: "r``{gy}=r``{gr}" "y=f`(gy)" "t=f`(gr)" by auto + from ff gyG grG B(2,3) have "yH" "tH" using apply_type by simp_all with equiv(G,r) grG r``{gy}=r``{gr} have "gy,grr" using same_image_equiv by simp - with assms(4) ff have + with assms(4) ff have "f`(P`gy,GroupInv(G,P)`(gr)) = TheNeutralElement(H,F)" unfolding QuotientGroupRel_def using func1_1_L15 by simp - with assms(1-4) B(2,3) gyG grG yH tH have "y=t" + with assms(1-4) B(2,3) gyG grG yH tH have "y=t" using image_inv group0.inverse_in_group group0.group0_2_L11A unfolding group0_def Homomor_def IsMorphism_def by auto } hence "x y. x,y ?𝔣 (z. x,z?𝔣 y=z)" by auto - ultimately have ff_fun: "?𝔣:G//rf``(G)" unfolding Pi_def function_def + ultimately have ff_fun: "?𝔣:G//rf``(G)" unfolding Pi_def function_def by auto - { fix a1 a2 assume AS: "a1G//r" "a2G//r" - then obtain g1 g2 where "g1G" "g2G" and a: "a1=r``{g1}" "a2=r``{g2}" + { fix a1 a2 assume AS: "a1G//r" "a2G//r" + then obtain g1 g2 where "g1G" "g2G" and a: "a1=r``{g1}" "a2=r``{g2}" unfolding quotient_def by auto - with assms equiv(G,r) have "𝒫`a1,a2,f`(P`g1,g2) ?𝔣" + with assms equiv(G,r) have "𝒫`a1,a2,f`(P`g1,g2) ?𝔣" using Group_ZF_2_4_L5A kernel_normal_sub group0.Group_ZF_2_2_L2 group0.group_op_closed unfolding QuotientGroupOp_def group0_def by auto - with ff_fun have eq: "?𝔣`(𝒫`a1,a2) = f`(P`g1,g2)" using apply_equality + with ff_fun have eq: "?𝔣`(𝒫`a1,a2) = f`(P`g1,g2)" using apply_equality by simp - from g1G g2G a have "a1,f`(g1) ?𝔣" and "a2,f`(g2) ?𝔣" by auto - with assms(1,2,3) ff_fun g1G g2G eq have "F`?𝔣`(a1),?𝔣`(a2) = ?𝔣`(𝒫`a1,a2)" + from g1G g2G a have "a1,f`(g1) ?𝔣" and "a2,f`(g2) ?𝔣" by auto + with assms(1,2,3) ff_fun g1G g2G eq have "F`?𝔣`(a1),?𝔣`(a2) = ?𝔣`(𝒫`a1,a2)" using apply_equality unfolding Homomor_def IsMorphism_def by simp - moreover from AS ff_fun have "?𝔣`(a1) f``(G)" "?𝔣`(a2) f``(G)" + moreover from AS ff_fun have "?𝔣`(a1) f``(G)" "?𝔣`(a2) f``(G)" using apply_type by auto ultimately have "restrict(F,f``(G)×f``(G))`?𝔣`(a1),?𝔣`(a2) = ?𝔣`(𝒫`a1,a2)" by simp } hence r: "a1G//r. a2G//r. restrict(F,f``(G)×f``(G))`?𝔣`(a1),?𝔣`(a2) = ?𝔣`(𝒫`a1,a2)" by simp - with ff_fun have HOM: "Homomor(?𝔣,G//r,𝒫,f``(G),restrict(F,(f``(G))×(f``(G))))" + with ff_fun have HOM: "Homomor(?𝔣,G//r,𝒫,f``(G),restrict(F,(f``(G))×(f``(G))))" unfolding Homomor_def IsMorphism_def by simp - from assms have G: "IsAgroup(G//r,𝒫)" and - H: "IsAgroup(f``(G), restrict(F,f``(G)×f``(G)))" + from assms have G: "IsAgroup(G//r,𝒫)" and + H: "IsAgroup(f``(G), restrict(F,f``(G)×f``(G)))" using Group_ZF_2_4_T1 kernel_normal_sub image_group unfolding IsAsubgroup_def by simp_all - { fix b1 b2 assume AS: "?𝔣`(b1) = ?𝔣`(b2)" "b1G//r" "b2G//r" - from G AS(3) have invb2: "GroupInv(G//r,𝒫)`(b2)G//r" + { fix b1 b2 assume AS: "?𝔣`(b1) = ?𝔣`(b2)" "b1G//r" "b2G//r" + from G AS(3) have invb2: "GroupInv(G//r,𝒫)`(b2)G//r" using group0.inverse_in_group unfolding group0_def by simp - with G AS(2) have I: "𝒫`b1,GroupInv(G//r,𝒫)`(b2)G//r" + with G AS(2) have I: "𝒫`b1,GroupInv(G//r,𝒫)`(b2)G//r" using group0.group_op_closed unfolding group0_def by auto - then obtain g where "gG" and gg: "𝒫`b1,GroupInv(G//r,𝒫)`(b2)=r``{g}" + then obtain g where "gG" and gg: "𝒫`b1,GroupInv(G//r,𝒫)`(b2)=r``{g}" unfolding quotient_def by auto from gG have "r``{g},f`(g) ?𝔣" by blast - with ff_fun gg have E: "?𝔣`(𝒫`b1,GroupInv(G//r,𝒫)`(b2)) = f`(g)" + with ff_fun gg have E: "?𝔣`(𝒫`b1,GroupInv(G//r,𝒫)`(b2)) = f`(g)" using apply_equality by simp - from ff_fun invb2 have pp: "?𝔣`(GroupInv(G//r,𝒫)`(b2))f``(G)" + from ff_fun invb2 have pp: "?𝔣`(GroupInv(G//r,𝒫)`(b2))f``(G)" using apply_type by simp - from ff_fun AS(2,3) have fff: "?𝔣`(b1) f``(G)" "?𝔣`(b2) f``(G)" + from ff_fun AS(2,3) have fff: "?𝔣`(b1) f``(G)" "?𝔣`(b2) f``(G)" using apply_type by simp_all - from fff(1) pp have - EE: "F`?𝔣`(b1),?𝔣`(GroupInv(G//r,𝒫)`(b2))= + from fff(1) pp have + EE: "F`?𝔣`(b1),?𝔣`(GroupInv(G//r,𝒫)`(b2))= restrict(F,f``(G)×f``(G))`?𝔣`(b1),?𝔣`(GroupInv(G//r,𝒫)`(b2))" by auto - from ff have "f``(G) H" using func1_1_L6(2) by simp - with fff have "?𝔣`(b1)H" "?𝔣`(b2)H" by auto - with assms(1-4) G H HOM ff_fun AS(1,3) fff(2) EE have + from ff have "f``(G) H" using func1_1_L6(2) by simp + with fff have "?𝔣`(b1)H" "?𝔣`(b2)H" by auto + with assms(1-4) G H HOM ff_fun AS(1,3) fff(2) EE have "TheNeutralElement(H,F) = restrict(F,f``(G)×f``(G))`?𝔣`(b1),?𝔣`(GroupInv(G//r,𝒫)`(b2))" using group0.group0_2_L6(1) restrict image_inv group0.group0_3_T1 image_group unfolding group0_def by simp - also from G H HOM AS(2,3) E have " = f`(g)" + also from G H HOM AS(2,3) E have " = f`(g)" using group0.inverse_in_group unfolding group0_def IsMorphism_def Homomor_def by simp finally have "TheNeutralElement(H,F) = f`(g)" by simp - with ff gG have "gf-``{TheNeutralElement(H,F)}" using func1_1_L15 + with ff gG have "gf-``{TheNeutralElement(H,F)}" using func1_1_L15 by simp - with assms gG gg have + with assms gG gg have "𝒫`b1,GroupInv(G//r,𝒫)`(b2) = TheNeutralElement(G//r,𝒫)" using group0.Group_ZF_2_4_L5E kernel_normal_sub unfolding group0_def by simp - with AS(2,3) G have "b1=b2" using group0.group0_2_L11A unfolding group0_def + with AS(2,3) G have "b1=b2" using group0.group0_2_L11A unfolding group0_def by auto - } with ff_fun have "?𝔣 inj(G//r,f``(G))" unfolding inj_def by blast + } with ff_fun have "?𝔣 inj(G//r,f``(G))" unfolding inj_def by blast moreover { fix m assume "m f``(G)" - with ff obtain g where "gG" "m=f`(g)" using func_imagedef by auto + with ff obtain g where "gG" "m=f`(g)" using func_imagedef by auto hence "r``{g},m ?𝔣" by blast - with ff_fun have "?𝔣`(r``{g})=m" using apply_equality by auto + with ff_fun have "?𝔣`(r``{g})=m" using apply_equality by auto with gG have "AG//r. ?𝔣`(A) = m" unfolding quotient_def by auto } ultimately have "?𝔣 bij(G//r,f``G)" unfolding bij_def surj_def - using ff_fun by blast - with HOM show ?thesis by blast + using ff_fun by blast + with HOM show ?thesis by blast qed text‹The inverse of a bijective homomorphism is an homomorphism. Meaning that in the previous result, the homomorphism we found is an isomorphism.› -theorem bij_homomor: +theorem bij_homomor: assumes "fbij(G,H)" "IsAgroup(G,P)" "Homomor(f,G,P,H,F)" shows "Homomor(converse(f),H,F,G,P)" proof - { fix h1 h2 assume "h1H" "h2H" - with assms(1) obtain g1 g2 where - g1: "g1G" "f`(g1)=h1" and g2: "g2G" "f`(g2)=h2" + with assms(1) obtain g1 g2 where + g1: "g1G" "f`(g1)=h1" and g2: "g2G" "f`(g2)=h2" unfolding bij_def surj_def by blast - with assms(2,3) have + with assms(2,3) have "converse(f)`(f`(P`g1,g2)) = converse(f)`(F`h1,h2)" using homomor_eq by simp - with assms(1,2) g1 g2 have + with assms(1,2) g1 g2 have "P`converse(f)`(h1),converse(f)`(h2) = converse(f)`(F`h1,h2)" using left_inverse group0.group_op_closed unfolding group0_def bij_def by auto - } with assms(1) show ?thesis using bij_converse_bij bij_is_fun + } with assms(1) show ?thesis using bij_converse_bij bij_is_fun unfolding Homomor_def IsMorphism_def by simp qed @@ -666,11 +669,11 @@

Theory Group_ZF_5

is an alternative notation for function defined as a set of pairs, see lemma
lambda_fun_alt› in theory func1.thy›.›
-lemma (in group0) quotient_map: +lemma (in group0) quotient_map: assumes "IsAnormalSubgroup(G,P,H)" defines "r QuotientGroupRel(G,P,H)" and "q λxG. QuotientGroupRel(G,P,H)``{x}" shows "Homomor(q,G,P,G//r,QuotientGroupOp(G,P,H))" - using groupAssum assms group_op_closed lam_funtype lamE EquivClass_1_L10 + using groupAssum assms group_op_closed lam_funtype lamE EquivClass_1_L10 Group_ZF_2_4_L3 Group_ZF_2_4_L5A Group_ZF_2_4_T1 unfolding IsAnormalSubgroup_def QuotientGroupOp_def Homomor_def IsMorphism_def by simp diff --git a/docs/IsarMathLib/MetricSpace_ZF.html b/docs/IsarMathLib/MetricSpace_ZF.html index cd5004e..a6f5313 100644 --- a/docs/IsarMathLib/MetricSpace_ZF.html +++ b/docs/IsarMathLib/MetricSpace_ZF.html @@ -16,7 +16,7 @@

Theory MetricSpace_ZF

This file is a part of IsarMathLib - a library of formalized mathematics written for Isabelle/Isar. - Copyright (C) 2020,2021 Slawomir Kolodynski + Copyright (C) 2020 - 2024 Slawomir Kolodynski This program is free software; Redistribution and use in source and binary forms, with or without modification, are permitted provided that the following conditions are met: @@ -42,7 +42,7 @@

Theory MetricSpace_ZF

section ‹ Metric spaces › -theory MetricSpace_ZF imports Topology_ZF_1 OrderedLoop_ZF Lattice_ZF +theory MetricSpace_ZF imports Topology_ZF_1 OrderedLoop_ZF Lattice_ZF UniformSpace_ZF begin text‹A metric space is a set on which a distance between points is defined as a function @@ -79,45 +79,62 @@

Theory MetricSpace_ZF

definition "Disk(X,d,r,c,R) {xX. d`c,x,R StrictVersion(r)}" +text‹We define metric topology› as consisting of unions of open disks.› + +definition + "MetricTopology(X,L,A,r,d) {𝒜. 𝒜 Pow(cX. {Disk(X,d,r,c,R). RPositiveSet(L,A,r)})}" + text‹Next we define notation for metric spaces. We will reuse the additive notation defined in the loop1› locale adding only the assumption about $d$ being a pseudometric and notation for a disk centered at $c$ with radius $R$. Since for many theorems it is sufficient to assume the pseudometric axioms we will - assume in this context that the sets $d,X,L,A,r$ form a pseudometric raher than a metric.› + assume in this context that the sets $d,X,L,A,r$ form a pseudometric raher than a metric. + In the pmetric_space› context $\tau$ denotes the topology defined by the metric $d$. + Analogously to the notation defined in the topology0› context int(A)›, + cl(A)›, ∂A› will denote the interior, closure and boundary of the set $A$ + with respect to the metric topology. ›
locale pmetric_space = loop1 + fixes d and X - assumes pmetricAssum: "IsApseudoMetric(d,X,L,A,r)" + assumes pmetricAssum: "IsApseudoMetric(d,X,L,A,r)" fixes disk defines disk_def [simp]: "disk(c,R) Disk(X,d,r,c,R)" - + fixes pmettop ("τ") + defines pmettop [simp]: "τ MetricTopology(X,L,A,r,d)" + fixes interior ("int") + defines interior_def [simp]: "int(D) Interior(D,τ)" + fixes cl + defines cl_def [simp]: "cl(D) Closure(D,τ)" + text‹ The next lemma shows the definition of the pseudometric in the notation used in the - metric_space› context.› + pmetric_space› context.› -lemma (in pmetric_space) pmetric_properties: shows +lemma (in pmetric_space) pmetric_properties: shows "d: X×X L+" "xX. d`x,x = 𝟬" "xX.yX. d`x,y = d`y,x" "xX.yX.zX. d`x,z \<lsq> d`x,y \<ra> d`y,z" - using pmetricAssum unfolding IsApseudoMetric_def by auto + using pmetricAssum unfolding IsApseudoMetric_def by auto -text‹The values of the metric are in the loop.› +text‹The values of the metric are in the in the nonnegative set of the loop, + hence in the loop.› -lemma (in pmetric_space) pmetric_loop_valued: assumes "xX" "yX" +lemma (in pmetric_space) pmetric_loop_valued: assumes "xX" "yX" shows "d`x,y L+" "d`x,y L" proof - - from assms show "d`x,y L+" using pmetric_properties(1) apply_funtype + from assms show "d`x,y L+" using pmetric_properties(1) apply_funtype by simp then show "d`x,y L" using Nonnegative_def by auto qed text‹The definition of the disk in the notation used in the pmetric_space› context:› -lemma (in pmetric_space) disk_definition: shows "disk(c,R) = {xX. d`c,x \<ls> R}" +lemma (in pmetric_space) disk_definition: shows "disk(c,R) = {xX. d`c,x \<ls> R}" proof - have "disk(c,R) = Disk(X,d,r,c,R)" by simp - then have "disk(c,R) = {xX. d`c,x,R StrictVersion(r)}" unfolding Disk_def by simp + then have "disk(c,R) = {xX. d`c,x,R StrictVersion(r)}" + unfolding Disk_def by simp moreover have "xX. d`c,x,R StrictVersion(r) d`c,x \<ls> R" using def_of_strict_ver by simp ultimately show ?thesis by auto @@ -125,49 +142,74 @@

Theory MetricSpace_ZF

text‹If the radius is positive then the center is in disk.› -lemma (in pmetric_space) center_in_disk: assumes "cX" and "RL+" shows "c disk(c,R)" - using pmetricAssum assms IsApseudoMetric_def PositiveSet_def disk_definition by simp +lemma (in pmetric_space) center_in_disk: assumes "cX" and "RL+" shows "c disk(c,R)" + using pmetricAssum assms IsApseudoMetric_def PositiveSet_def disk_definition by simp -text‹A technical lemma that allows us to shorten some proofs: › +text‹A technical lemma that allows us to shorten some proofs: if $c$ is an element + of $X$ and $x$ is in disk with center $c$ and radius $R$ then $R$ is a positive element of + $L$ and $-d(x,y)+R$ is in the set of positive elements of the loop. › -lemma (in pmetric_space) radius_in_loop: assumes "cX" and "x disk(c,R)" +lemma (in pmetric_space) radius_in_loop: assumes "cX" and "x disk(c,R)" shows "RL" "𝟬\<ls>R" "RL+" "(\<rm>d`c,x \<ad> R) L+" proof - - from assms(2) have "xX" and "d`c,x \<ls> R" using disk_definition by auto - with assms(1) show "𝟬\<ls>R" using pmetric_properties(1) apply_funtype + from assms(2) have "xX" and "d`c,x \<ls> R" using disk_definition by auto + with assms(1) show "𝟬\<ls>R" using pmetric_properties(1) apply_funtype nonneg_definition loop_strict_ord_trans by blast then show "RL" and "RL+" using posset_definition PositiveSet_def by auto from d`c,x \<ls> R show "(\<rm>d`c,x \<ad> R) L+" using ls_other_side(2) by simp qed -text‹If a point $x$ is inside a disk $B$ and $m\leq R-d(c,x)$ then the disk centered +text‹If a point $x$ is inside a disk $B$ and $m\leq -d\langle c,x\rangle + R$ then the disk centered at the point $x$ and with radius $m$ is contained in the disk $B$. › -lemma (in pmetric_space) disk_in_disk: +lemma (in pmetric_space) disk_in_disk: assumes "cX" and "x disk(c,R)" and "m \<lsq> (\<rm>d`c,x \<ad> R)" shows "disk(x,m) disk(c,R)" proof fix y assume "y disk(x,m)" - then have "d`x,y \<ls> m" using disk_definition by simp - from assms(1,2) y disk(x,m) have "RL" "xX" "yX" - using radius_in_loop(1) disk_definition by auto - with assms(1) have "d`c,y \<lsq> d`c,x \<ra> d`x,y" using pmetric_properties(4) by simp - from assms(1) xX have "d`c,x L" - using pmetric_properties(1) apply_funtype nonneg_subset by auto - with d`x,y \<ls> m assms(3) have "d`c,x \<ra> d`x,y \<ls> d`c,x \<ra> (\<rm>d`c,x \<ad> R)" + then have "d`x,y \<ls> m" using disk_definition by simp + from assms(1,2) y disk(x,m) have "RL" "xX" "yX" + using radius_in_loop(1) disk_definition by auto + with assms(1) have "d`c,y \<lsq> d`c,x \<ra> d`x,y" using pmetric_properties(4) by simp + from assms(1) xX have "d`c,x L" + using pmetric_properties(1) apply_funtype nonneg_subset by auto + with d`x,y \<ls> m assms(3) have "d`c,x \<ra> d`x,y \<ls> d`c,x \<ra> (\<rm>d`c,x \<ad> R)" using loop_strict_ord_trans1 strict_ord_trans_inv(2) by blast with d`c,x L RL d`c,y \<lsq> d`c,x \<ra> d`x,y yX show "y disk(c,R)" - using lrdiv_props(6) loop_strict_ord_trans disk_definition by simp + using lrdiv_props(6) loop_strict_ord_trans disk_definition by simp +qed + +text‹A special case of disk_in_disk› where we set $m = -d\langle c,x\rangle + R$: + if $x$ is an element of a disk with center $c\in X$ + and radius $R$ then this disk contains the disk centered at $x$ and with radius + $-d\langle c,x\rangle + R$. › + +lemma (in pmetric_space) disk_in_disk1: + assumes "cX" and "x disk(c,R)" + shows "disk(x,\<rm>d`c,x \<ad> R) disk(c,R)" +proof - + from assms(2) have "RL" and "d`c,x L" + using disk_definition less_members by auto + with assms show ?thesis using left_right_sub_closed(1) loop_ord_refl disk_in_disk + by simp qed -text‹ If we assume that the order on the group makes the positive set a meet semi-lattice (i.e. - every two-element subset of $G_+$ has a greatest lower bound) then +text‹Assuming that two disks have the same center, closed disk with smaller radius + in contained in the (open) disk with a larger radius. › + +lemma (in pmetric_space) disk_radius_strict_mono: + assumes "r1 \<ls> r2" + shows "{yX. d`x,y \<lsq> r1} disk(x,r2)" + using assms loop_strict_ord_trans disk_definition by auto + +text‹ If we assume that the loop's order relation down-directs $L_+$ then the collection of disks centered at points of the space and with radii in the positive set - of the group satisfies the base condition. The meet semi-lattice assumption can be weakened - to "each two-element subset of $G_+$ has a lower bound in $G_+$", but we don't do that here. › + of the loop satisfies the base condition. The property that an order relation "down-directs" + a set is defined in Order_ZF› and means that every two-element subset of the set + has a lower bound in that set. › -lemma (in pmetric_space) disks_form_base: +lemma (in pmetric_space) disks_form_base: assumes "r {down-directs} L+" defines "B cX. {disk(c,R). RL+}" shows "B {satisfies the base condition}" @@ -176,23 +218,23 @@

Theory MetricSpace_ZF

fix x assume "xUV" have "WB. xW WUV" proof - - from assms(2) UB VB obtain cU cV RU RV + from assms(2) UB VB obtain cU cV RU RV where "cU X" "RU L+" "cV X" "RV L+" "U = disk(cU,RU)" "V = disk(cV,RV)" by auto with xUV have "x disk(cU,RU)" and "x disk(cV,RV)" by auto - then have "xX" "d`cU,x \<ls> RU" "d`cV,x \<ls> RV" using disk_definition by auto + then have "xX" "d`cU,x \<ls> RU" "d`cV,x \<ls> RV" using disk_definition by auto let ?mU = "\<rm> d`cU,x \<ad> RU" let ?mV = "\<rm> d`cV,x \<ad> RV" from cUX xdisk(cU,RU) cVX xdisk(cV,RV) have "?mUL+" and "?mVL+" - using radius_in_loop(4) by auto - with assms(1) obtain m where "m L+" "m \<lsq> ?mU" "m \<lsq> ?mV" + using radius_in_loop(4) by auto + with assms(1) obtain m where "m L+" "m \<lsq> ?mU" "m \<lsq> ?mV" unfolding DownDirects_def by auto let ?W = "disk(x,m)" from m L+ m \<lsq> ?mU m \<lsq> ?mV cU X x disk(cU,RU) cV X x disk(cV,RV) U = disk(cU,RU) V = disk(cV,RV) - have "?W UV" using disk_in_disk by blast - moreover from assms(2) xX m L+ have "?W B" and "x?W" using center_in_disk + have "?W UV" using disk_in_disk by blast + moreover from assms(2) xX m L+ have "?W B" and "x?W" using center_in_disk by auto ultimately show ?thesis by auto qed @@ -200,138 +242,445 @@

Theory MetricSpace_ZF

qed text‹Disks centered at points farther away than the sum of radii do not overlap. › -lemma (in pmetric_space) far_disks: + +lemma (in pmetric_space) far_disks: assumes "xX" "yX" "rx\<ra>ry \<lsq> d`x,y" - shows "disk(x,rx)disk(y,ry) = 0" + shows "disk(x,rx)disk(y,ry) = " proof - - { assume "disk(x,rx)disk(y,ry) 0" + { assume "disk(x,rx)disk(y,ry) " then obtain z where "z disk(x,rx)disk(y,ry)" by auto then have "zX" and "d`x,z \<ra> d`y,z \<ls> rx\<ra>ry" - using disk_definition add_ineq_strict by auto - moreover from assms(1,2) zX have "d`x,y \<lsq> d`x,z \<ra> d`y,z" - using pmetric_properties(3,4) by auto + using disk_definition add_ineq_strict by auto + moreover from assms(1,2) zX have "d`x,y \<lsq> d`x,z \<ra> d`y,z" + using pmetric_properties(3,4) by auto ultimately have "d`x,y \<ls> rx\<ra>ry" using loop_strict_ord_trans by simp - with assms(3) have False using loop_strict_ord_trans by auto + with assms(3) have False using loop_strict_ord_trans by auto } thus ?thesis by auto qed text‹ If we have a loop element that is smaller than the distance between two points, then we can separate these points with disks.› -lemma (in pmetric_space) disjoint_disks: +lemma (in pmetric_space) disjoint_disks: assumes "xX" "yX" "rx\<ls>d`x,y" shows "(\<rm>rx\<ad>(d`x,y)) L+" and "disk(x,rx)disk(y,\<rm>rx\<ad>(d`x,y)) = 0" proof - - from assms(3) show "(\<rm>rx\<ad>(d`x,y)) L+" + from assms(3) show "(\<rm>rx\<ad>(d`x,y)) L+" using ls_other_side posset_definition1 by simp - from assms(1,2,3) have "rxL" and "d`x,y L" - using less_members(1) pmetric_loop_valued(2) by auto + from assms(1,2,3) have "rxL" and "d`x,y L" + using less_members(1) pmetric_loop_valued(2) by auto then have "rx\<ra>(\<rm>rx\<ad>(d`x,y)) = d`x,y" using lrdiv_props(6) by simp - with assms(1,2) d`x,y L show "disk(x,rx)disk(y,\<rm>rx\<ad>(d`x,y)) = 0" - using loop_ord_refl far_disks by simp + with assms(1,2) d`x,y L show "disk(x,rx)disk(y,\<rm>rx\<ad>(d`x,y)) = 0" + using loop_ord_refl far_disks by simp +qed + +text‹The definition of metric topology written in notation of pmetric_space› context:› + +lemma (in pmetric_space) metric_top_def_alt: + defines "B cX. {disk(c,R). RL+}" + shows "τ = {A. A Pow(B)}" +proof - + from assms have "MetricTopology(X,L,A,r,d) = {A. A Pow(B)}" + unfolding MetricTopology_def by simp + thus ?thesis by simp qed -text‹Unions of disks form a topology, hence (pseudo)metric spaces are topological spaces. › +text‹If the order of the loop down-directs its set of positive elements + then the metric topology defined as collection of unions of (open) disks is indeed a topology. + Recall that in the pmetric_space› context $\tau$ denotes the metric topology. › + +theorem (in pmetric_space) pmetric_is_top: + assumes "r {down-directs} L+" + shows "τ {is a topology}" + using assms disks_form_base Top_1_2_T1 metric_top_def_alt by simp + +text‹If $r$ down-directs $L_+$ then the collection of open disks is a base for + the metric topology.› -theorem (in pmetric_space) pmetric_is_top: +theorem (in pmetric_space) disks_are_base: assumes "r {down-directs} L+" - defines "B cX. {disk(c,R). RL+}" - defines "T {A. A Pow(B)}" - shows "T {is a topology}" "B {is a base for} T" "T = X" + defines "B cX. {disk(c,R). RL+}" + shows "B {is a base for} τ" + using assms disks_form_base Top_1_2_T1 metric_top_def_alt by simp + +text‹If $r$ down-directs $L_+$ then $X$ is the carrier of metric topology.› + +theorem (in pmetric_space) metric_top_carrier: + assumes "r {down-directs} L+" shows "τ = X" proof - - from assms show "T {is a topology}" "B {is a base for} T" - using disks_form_base Top_1_2_T1 by auto - then have "T = B" using Top_1_2_L5 by simp - moreover have "B = X" + let ?B = "cX. {disk(c,R). RL+}" + from assms have "τ = ?B" + using disks_are_base Top_1_2_L5 by simp + moreover have "?B = X" proof - from assms(2) show "B X" using disk_definition by auto - { fix x assume "xX" - from assms(1) obtain R where "RL+" unfolding DownDirects_def by blast - with assms(2) xX have "x B" using center_in_disk by auto - } thus "X B" by auto + show "?B X" using disk_definition by auto + from assms show "X ?B" unfolding DownDirects_def using center_in_disk + by blast qed - ultimately show "T = X" by simp + ultimately show "τ = X" by simp qed +text‹Under the assumption that $r$ down-directs $L_+$ the propositions proven + in the topology0› context can be used in the pmetric_space› context.› + +lemma (in pmetric_space) topology0_valid_in_pmetric_space: + assumes "r {down-directs} L+" + shows "topology0(τ)" + using assms pmetric_is_top unfolding topology0_def by simp + +text‹If $r$ down-directs $L_+$ then disks are open in the metric topology.› + +lemma (in pmetric_space) disks_open: + assumes "cX" "RL+" "r {down-directs} L+" + shows "disk(c,R) τ" + using assms base_sets_open disks_are_base(1) pmetric_is_top + by blast + text‹To define the metric_space› locale we take the pmetric_space› and add the assumption of identity of indiscernibles.› locale metric_space = pmetric_space + - assumes ident_indisc: "xX. yX. d`x,y=𝟬 x=y" + assumes ident_indisc: "xX. yX. d`x,y=𝟬 x=y" text‹In the metric_space› locale $d$ is a metric.› lemma (in metric_space) d_metric: shows "IsAmetric(d,X,L,A,r)" - using pmetricAssum ident_indisc unfolding IsAmetric_def by simp + using pmetricAssum ident_indisc unfolding IsAmetric_def by simp text‹Distance of different points is greater than zero. › -lemma (in metric_space) dist_pos: assumes "xX" "yX" "xy" +lemma (in metric_space) dist_pos: assumes "xX" "yX" "xy" shows "𝟬\<ls>d`x,y" "d`x,y L+" proof - - from assms(1,2) have "d`x,y L+" - using pmetric_properties(1) apply_funtype by simp + from assms(1,2) have "d`x,y L+" + using pmetric_properties(1) apply_funtype by simp then have "𝟬 \<lsq> d`x,y" using Nonnegative_def by auto - with assms show "d`x,y L+" and "𝟬\<ls>d`x,y" - using ident_indisc posset_definition posset_definition1 by auto + with assms show "d`x,y L+" and "𝟬\<ls>d`x,y" + using ident_indisc posset_definition posset_definition1 by auto qed -text‹An ordered loop valued metric space is $T_2$ (i.e. Hausdorff).› +text‹If $r$ down-directs $L_+$ then the ordered loop valued metric space is $T_2$ (i.e. Hausdorff).› -theorem (in metric_space) metric_space_T2: - assumes "r {down-directs} L+" - defines "B cX. {disk(c,R). RL+}" - defines "T {A. A Pow(B)}" - shows "T {is T2}" +theorem (in metric_space) metric_space_T2: + assumes "r {down-directs} L+" + shows "τ {is T2}" proof - - { fix x y assume "xT" "yT" "xy" - from assms have "BT" using pmetric_is_top(2) base_sets_open by auto - moreover have "UB. VB. xU yV UV = 0" + let ?B = "cX. {disk(c,R). RL+}" + { fix x y assume "xτ" "yτ" "xy" + from assms have "?Bτ" using metric_top_def_alt by auto + have "U?B. V?B. xU yV UV = " proof - - let ?R = "d`x,y" - from assms have "T = X" using pmetric_is_top(3) by simp - with xT yT have "xX" "yX" by auto - with xy have "?RL+" using dist_pos by simp - with assms(2) xX yX have "disk(x,?R) B" and "disk(y,?R) B" + let ?R = "d`x,y" + from assms have "τ = X" using metric_top_carrier by simp + with xτ have "xX" by blast + from τ = X yτ have "yX" by blast + with xy xX have "?RL+" using dist_pos by simp + with xX yX have "disk(x,?R) ?B" and "disk(y,?R) ?B" by auto - { assume "disk(x,?R) disk(y,?R) = 0" - moreover from assms(2) xX yX ?RL+ have - "disk(x,?R)B" "disk(y,?R)B" "xdisk(x,?R)" "ydisk(y,?R)" - using center_in_disk by auto - ultimately have "UB. VB. xU yV UV = 0" by auto + { assume "disk(x,?R) disk(y,?R) = " + moreover from xX yX ?RL+ have + "disk(x,?R)?B" "disk(y,?R)?B" "xdisk(x,?R)" "ydisk(y,?R)" + using center_in_disk by auto + ultimately have "U?B. V?B. xU yV UV = 0" by blast } moreover { assume "disk(x,?R) disk(y,?R) 0" then obtain z where "z disk(x,?R)" and "z disk(y,?R)" by auto - then have "d`x,z \<ls> ?R" using disk_definition by simp + then have "d`x,z \<ls> ?R" using disk_definition by simp then have "𝟬 \<ls> \<rm>d`x,z\<ad>?R" using ls_other_side(1) by simp let ?r = "\<rm>d`x,z\<ad>?R" have "?r\<ls>?R" proof - from z disk(y,?R) xX yX have "zX" "xz" - using disk_definition pmetric_properties(3) by auto + using disk_definition pmetric_properties(3) by auto with xX yX zX show ?thesis - using pmetric_loop_valued dist_pos(1) subtract_pos(2) by simp + using pmetric_loop_valued dist_pos(1) add_subtract_pos(2) by simp qed with xX yX have "disk(x,?r)disk(y,\<rm>?r\<ad>?R) = 0" - by (rule disjoint_disks) + by (rule disjoint_disks) moreover from 𝟬\<ls>?r ?r\<ls>?R have "?rL+" "(\<rm>?r\<ad>?R) L+" using ls_other_side posset_definition1 by auto - with assms(2) xX yX have - "disk(x,?r)B" "disk(y,\<rm>?r\<ad>(d`x,y))B" and + with xX yX have + "disk(x,?r)?B" "disk(y,\<rm>?r\<ad>(d`x,y))?B" and "xdisk(x,?r)" "ydisk(y,\<rm>?r\<ad>(d`x,y))" - using center_in_disk by auto - ultimately have "UB. VB. xU yV UV = 0" by auto + using center_in_disk by auto + ultimately have "U?B. V?B. xU yV UV = 0" by blast } ultimately show ?thesis by auto qed - ultimately have "UT. VT. xU yV UV = 0" by auto + with ?Bτ have "Uτ. Vτ. xU yV UV = " by (rule exist2_subset) } then show ?thesis unfolding isT2_def by simp qed +subsection‹Uniform structures on metric spaces› + +text‹Each pseudometric space with pseudometric $d:X\times X\rightarrow L$ + supports a natural uniform structure, defined as supersets of the collection + of inverse images $U_c = d^{-1}([0,c])$, where $c>0$. › + +text‹In the following definition $X$ is the underlying space, $L$ is the loop (carrier), + $A$ is the loop operation, $r$ is an order relation compatible with $A$, + and $d$ is a pseudometric on $X$, valued in the ordered loop $L$. + With this we define the uniform gauge as the collection of inverse images + of the closed intervals $[0,c]$ as $c$ varies of the set of positive elements of $L$.› + +definition + "UniformGauge(X,L,A,r,d) {d-``({bNonnegative(L,A,r). b,c r}). cPositiveSet(L,A,r)}" + +text‹In the pmetric_space› context we will write UniformGauge(X,L,A,r,d)› as 𝔅›. › + +abbreviation (in pmetric_space) gauge ("𝔅") where "𝔅 UniformGauge(X,L,A,r,d)" + +text‹In notation defined in the pmetric_space› context we can write the uniform gauge + as $\{ d^{-1}(\{c\in L^+: c\leq b\}: b \in L_+ \}$. › + +lemma (in pmetric_space) uniform_gauge_def_alt: + shows "𝔅 = {d-``({cL+. c\<lsq>b}). bL+}" + unfolding UniformGauge_def by simp + +text‹Members of the uniform gauge are subsets of $X\times X$ i.e. relations on $X$. › + +lemma (in pmetric_space) uniform_gauge_relations: + assumes "B𝔅" shows "BX×X" + using assms uniform_gauge_def_alt pmetric_properties(1) func1_1_L3 + by force + +text‹If the distance between two points of $X$ is less or equal $b$, then + this pair of points is in $d^{-1}([0,b])$. › + +lemma (in pmetric_space) gauge_members: + assumes "xX" "yX" "d`x,y \<lsq> b" + shows "x,y d-``({cL+. c\<lsq>b})" + using assms pmetric_properties(1) apply_funtype func1_1_L15 + by simp + +text‹Suppose $b\in L^+$ (i.e. b is an element of the loop that is greater than the neutral element) + and $x\in X$. Then the image of the singleton set $\{ x\}$ by the relation + $B=\{ d^{-1}(\{c\in L^+: c\leq b\}$ is the set $\{ y\in X:d\langle x,y\rangle \leq b\}$, + i.e. the closed disk with center $x$ and radius $b$. Hence the the image $B\{ x\}$ contains + the open disk with center $x$ and radius $b$. › + +lemma (in pmetric_space) disk_in_gauge: + assumes "bL+" "xX" + defines "B d-``({cL+. c\<lsq>b})" + shows "B``{x} = {yX. d`x,y \<lsq> b}" and "disk(x,b) B``{x}" +proof - + from assms(1,3) have "BX×X" + using uniform_gauge_def_alt uniform_gauge_relations by auto + with assms(2,3) show "B``{x} = {yX. d`x,y \<lsq> b}" + using pmetric_properties(1) func1_1_L15 by force + then show "disk(x,b) B``{x}" using disk_definition by auto +qed + +text‹Gauges corresponding to larger elements of the loop are larger. › + +lemma (in pmetric_space) uniform_gauge_mono: + assumes "b1\<lsq>b2" shows "d-``({cL+. c\<lsq>b1}) d-``({cL+. c\<lsq>b2})" + using ordLoopAssum assms vimage_mono1 + unfolding IsAnOrdLoop_def IsPartOrder_def trans_def by auto + +text‹For any two sets of the form $d^{-1}([0,b])$ we can find a third one that is contained + in both. › + +lemma (in pmetric_space) gauge_1st_cond: + assumes "r {down-directs} L+" "B1𝔅" "B2𝔅" + shows "B3𝔅. B3B1B2" +proof - + from assms(2,3) obtain b1 b2 where "b1L+" "b2L+" and + I: "B1 = d-``({cL+. c\<lsq>b1})" "B2 = d-``({cL+. c\<lsq>b2})" + using uniform_gauge_def_alt by auto + from assms(1) b1L+ b2L+ obtain b3 where "b3L+" "b3\<lsq>b1" "b3\<lsq>b2" + unfolding DownDirects_def by auto + from I b3\<lsq>b1 b3\<lsq>b2 have "d-``({cL+. c\<lsq>b3}) B1B2" + using uniform_gauge_mono by blast + with b3L+ show ?thesis using uniform_gauge_def_alt + by auto +qed + +text‹Sets of the form $d^{-1}([0,b])$ contain the diagonal. › + +lemma (in pmetric_space) gauge_2nd_cond: assumes "B𝔅" shows "id(X)B" +proof + fix p assume "pid(X)" + then obtain x where "xX" and "p=x,x" by auto + then have "pX×X" and "d`(p) = 𝟬" using pmetric_properties(2) by simp_all + from assms obtain b where "bL+" and "B = d-``({cL+. c\<lsq>b})" + using uniform_gauge_def_alt by auto + with pX×X d`(p) = 𝟬 show "pB" + using posset_definition1 loop_zero_nonneg pmetric_properties(1) func1_1_L15 + by simp +qed + +text‹Sets of the form $d^{-1}([0,b])$ are symmetric.› + +lemma (in pmetric_space) gauge_symmetric: + assumes "B𝔅" shows "B = converse(B)" +proof - + from assms obtain b where "B = d-``({cL+. c\<lsq>b})" + using uniform_gauge_def_alt by auto + with pmetricAssum show ?thesis unfolding IsApseudoMetric_def + using symm_vimage_symm by auto +qed + +text‹A set of the form $d^{-1}([0,b])$ contains a symmetric set of this form.› + +corollary (in pmetric_space) gauge_3rd_cond: + assumes "B1𝔅" shows "B2𝔅. B2 converse(B1)" + using assms gauge_symmetric by auto + +text‹The collection of sets of the form $d^{-1}([0,b])$ for $b\in L_+$ + is contained of the powerset of $X\times X$.› + +lemma (in pmetric_space) gauge_5thCond: shows "𝔅Pow(X×X)" + using uniform_gauge_def_alt pmetric_properties(1) func1_1_L3 by force + +text‹If the set of positive values is non-empty, then there are sets + of the form $d^{-1}([0,b])$ for $b>0$.› + +lemma (in pmetric_space) gauge_6thCond: + assumes "L+" shows "𝔅" using assms uniform_gauge_def_alt by simp + +text‹The remaining 4th condition for the sets of the form $d^{-1}([0,b])$ + to be a uniform base (or a fundamental system of entourages) cannot be proven + without additional assumptions in the context of ordered loop valued metrics. + To see that consider the example + of natural numbers with the metric $d\langle x,y \rangle = |x-y|$, where we think + of $d$ as valued in the nonnegative set of ordered group of integers. + Now take the set $B_1 = d^{-1}([0,1]) = d^{-1}(\{ 0,1\} )$. Then the set $B_1 \circ B_1$ + is strictly larger than $B_1$, but there is no smaller set $B_2$ we can take so that + $B_2 \circ B_2 \subseteq B_1$. + One condition that is sufficient is that for every $b_1 >0$ there is a $b_2 >0$ + such that $b_2 + b_2 \leq b_1 $. I have not found a standard name for this property, for now + we will use the name IsHalfable›. › + +definition + "IsHalfable(L,A,r) b1PositiveSet(L,A,r). b2PositiveSet(L,A,r). A`b2,b2,b1 r" + +text‹The property of halfability written in the notation used in the pmetric_space› context.› + +lemma (in pmetric_space) is_halfable_def_alt: + assumes "IsHalfable(L,A,r)" "b1L+" + shows "b2L+. b2\<ra>b2 \<lsq> b1" + using assms unfolding IsHalfable_def by simp + +text‹If the loop order is halfable then for every set $B_1$ of the form $d^{-1}([0,b_1])$ + for some $b_1>0$ we can find another one $B_2 = d^{-1}([0,b_2])$ such that $B_2$ + composed with itself is contained in $B_1$.› + +lemma (in pmetric_space) gauge_4thCond: + assumes "IsHalfable(L,A,r)" "B1𝔅" shows "B2𝔅.B2𝔅. B2 O B2 B1" +proof - + from assms(2) obtain b1 where "b1L+" and "B1 = d-``({cL+. c\<lsq>b1})" + using uniform_gauge_def_alt by auto + from assms(1) b1L+ obtain b2 where "b2L+" and "b2\<ra>b2 \<lsq> b1" + using is_halfable_def_alt by auto + let ?B2 = "d-``({cL+. c\<lsq>b2})" + from b2L+ have "?B2𝔅" unfolding UniformGauge_def by auto + { fix p assume "p ?B2 O ?B2" + with ?B2𝔅 obtain x y where "xX" "yX" and "p=x,y" + using gauge_5thCond by blast + from p ?B2 O ?B2 p=x,y obtain z where + "x,z ?B2" and "z,y ?B2" + using rel_compdef by auto + with ?B2𝔅 have "zX" using gauge_5thCond by auto + from x,z ?B2 z,y ?B2 have "d`x,z \<ra> d`z,y \<lsq> b2\<ra> b2" + using pmetric_properties(1) func1_1_L15 add_ineq by simp + with b2\<ra>b2 \<lsq> b1 have "d`x,z \<ra> d`z,y \<lsq> b1" + using loop_ord_trans by simp + with xX yX zX p=x,y B1 = d-``({cL+. c\<lsq>b1}) have "pB1" + using pmetric_properties(4) loop_ord_trans gauge_members by blast + } hence "?B2 O ?B2 B1" by auto + with ?B2𝔅 show ?thesis by auto +qed + +text‹If $X$ and $L_+$ are not empty, the order relation $r$ + down-directs $L_+$, and the loop order is halfable, then $\mathfrak{B}$ + (which in the pmetric_space› context is an abbreviation for + $\{ d^{-1}(\{c\in L^+: c\leq b\}: b \in L_+ \}$) + is a fundamental system of entourages, hence its supersets + form a uniformity on $X$ and hence those supersets define a topology on $X$.› + +theorem (in pmetric_space) metric_gauge_base: + assumes "X" "L+" "r {down-directs} L+" "IsHalfable(L,A,r)" + shows + "𝔅 {is a uniform base on} X" + "Supersets(X×X,𝔅) {is a uniformity on} X" + "UniformTopology(Supersets(X×X,𝔅),X) {is a topology}" + "UniformTopology(Supersets(X×X,𝔅),X) = X" + using assms gauge_1st_cond gauge_2nd_cond gauge_3rd_cond + gauge_4thCond gauge_5thCond gauge_6thCond uniformity_base_is_base + uniform_top_is_top + unfolding IsUniformityBaseOn_def by simp_all + +text‹At this point we know that a pseudometric induces two topologies: one consisting of unions + of open disks (the metric topology) and second one being the uniform topology derived + from the uniformity generated the fundamental system of entourages (the base uniformity) + of the sets of the form $d^{-1}([0,b])$ for $b>0$. + The next theorem states that if $X$ and $L_+$ are not empty, $r$ down-directs $L_+$, + and the loop order is halfable, then these two topologies are in fact the same. + Recall that in the pmetric_space› context $\tau$ denotes the metric topology. › + +theorem (in pmetric_space) metric_top_is_uniform_top: + assumes "X" "L+" "r {down-directs} L+" "IsHalfable(L,A,r)" + shows "τ = UniformTopology(Supersets(X×X,𝔅),X)" +proof + let = "Supersets(X×X,𝔅)" + from assms have " {is a uniformity on} X" using metric_gauge_base + by simp + let ?T = "UniformTopology(,X)" + { fix U assume "U?T" + then have "UPow(X)" and I: "xU. U{V``{x}. V}" + unfolding UniformTopology_def by auto + { fix x assume "xU" + with I obtain A where "A" and "U = A``{x}" + by auto + from xU U?T have "x?T" by auto + with assms have "xX" using metric_gauge_base(4) by simp + from A obtain B where "B𝔅" and "BA" + unfolding Supersets_def by auto + from B𝔅 obtain b where "bL+" and "B = d-``({cL+. c\<lsq>b})" + using uniform_gauge_def_alt by auto + with xX BA U = A``{x} have "disk(x,b) U" + using disk_in_gauge(2) by blast + with assms(3) xX bL+ have "Vτ. xV VU" + using disks_open center_in_disk by force + } with assms(3) have "Uτ" + using topology0_valid_in_pmetric_space topology0.open_neigh_open + by simp + } thus "?T τ" by auto + let ?𝒟 = "cX. {disk(c,R). RL+}" + { fix U assume "U ?𝒟" + then obtain c b where "cX" "bL+" "U = disk(c,b)" + by blast + { fix x assume "xU" + let ?b1 = "\<rm>d`c,x \<ad> b" + from xU cX U = disk(c,b) have + "xX" "xdisk(c,b)" "disk(x,?b1) U" "?b1 L+" + using disk_in_disk1 disk_definition radius_in_loop(4) by simp_all + with assms(4) obtain b2 where "b2L+" and "b2\<ra>b2 \<lsq> ?b1" + using is_halfable_def_alt by auto + let ?D = "{yX. d`x,y \<lsq> b2}" + from b2L+ b2\<ra>b2 \<lsq> ?b1 have "?D disk(x,?b1)" + using posset_definition1 positive_subset add_subtract_pos(3) + loop_strict_ord_trans1 disk_radius_strict_mono by blast + let ?B = "d-``({cL+. c\<lsq>b2})" + from b2L+ have "?B𝔅" using uniform_gauge_def_alt by auto + then have "?B" using uniform_gauge_relations superset_gen + by simp + from b2L+ xX ?D disk(x,?b1) disk(x,?b1) U have "?B``{x} U" + using disk_in_gauge(1) by auto + with ?B have "W. W``{x} U" by auto + } with U = disk(c,b) {is a uniformity on} X have "U ?T" + using disk_definition uniftop_def_alt1 by auto + } hence "?𝒟 ?T" by auto + with assms show "τ?T" + using disks_are_base(1) metric_gauge_base(3) base_smallest_top + by simp +qed + end diff --git a/docs/IsarMathLib/MetricSpace_ZF_1.html b/docs/IsarMathLib/MetricSpace_ZF_1.html index d4c5384..fc3c59e 100644 --- a/docs/IsarMathLib/MetricSpace_ZF_1.html +++ b/docs/IsarMathLib/MetricSpace_ZF_1.html @@ -58,79 +58,110 @@

Theory MetricSpace_ZF_1

positive numbers, absolute value etc. For metric spaces we reuse the notation defined there.›
text‹The pmetric_space1› locale extends the reals› locale, adding the carrier $X$ - of the metric space and the metric $\mathfrak{d}$ to the context, together with the assumption - that $\mathfrak{d}:X\times X \rightarrow \mathbb{R}$ is a pseudo metric. + of the metric space and the metric $\mathcal{d}$ to the context, together with the assumption + that $\mathcal{d}:X\times X \rightarrow \mathbb{R}^+$ is a pseudo metric. + We choose to denote the disk in $X$ with center $c$ and radius $r$ as ball(c,r)› + As in the pmetric_space› locale we define the $\tau$ to be the metric topology, i.e. + the topology induced by the (real valued) pseudometric $\mathcal{d}$. An alternative would be to define the pmetric_space1› as an extension of the pmetric_space1› context, but that is in turn an extension of the loop1› locale that defines notation for left and right division which which do not want in the context of real numbers. › locale pmetric_space1 = reals + fixes X and 𝖽 - assumes pmetricAssum: "IsApseudoMetric(𝖽,X,,Add,ROrd)" + assumes pmetricAssum: "IsApseudoMetric(𝖽,X,,Add,ROrd)" fixes ball defines ball_def [simp]: "ball(c,r) Disk(X,𝖽,ROrd,c,r)" + fixes pmettop ("τ") + defines pmettop [simp]: "τ MetricTopology(X,,Add,ROrd,𝖽)" + fixes interior ("int") + defines interior_def [simp]: "int(D) Interior(D,τ)" + fixes cl + defines cl_def [simp]: "cl(D) Closure(D,τ)" text‹The propositions proven in the pmetric_space› context defined in Metric_Space_ZF› theory are valid in the pmetric_space1› context. › -lemma (in pmetric_space1) pmetric_space_pmetric_space1_valid: +lemma (in pmetric_space1) pmetric_space_pmetric_space1_valid: shows "pmetric_space(,Add,ROrd,𝖽,X)" unfolding pmetric_space_def pmetric_space_axioms_def loop1_def - using pmetricAssum reals_loop by simp - -text‹It is convenient to have the collection of all open balls in given (p) metrics defined - as a separate notion.› - -definition (in pmetric_space1) Open_Balls - where "Open_Balls cX. {ball(c,r). r +}" - -text‹Topology on a metric space is defined as the collection of sets that are unions - of open balls of the (p)metric. › - -definition (in pmetric_space1) Metric_Topology - where "Metric_Topology {A. A Pow(Open_Balls)}" + using pmetricAssum reals_loop by simp + +text‹The context pmetric_space1› is a special case of context pmetric_space› + where the fixed objects in pmetric_space› map to (in the order defined in pmetric_space›) + the set of real numbers, real addition, the order relation on reals, + the strict order relation on reals, the set of non-negative reals and + the set of positive reals. The metrics $d$ maps to the real metrics + 𝖽›, the carrier of the metric space $X$ is still $X$, and the disk›s from pmetric_space› + are now called ball›s in pmetric_space1›. The notation for right and left division from + pmetric_space1› is not used in pmetric_space›. › + +sublocale pmetric_space1 < pmetric_space + "" Add ROrd "𝟬" realadd lesseq sless nonnegative positiveset + "λx y. LeftDiv(,Add)`x,y" + "λx y. RightDiv(,Add)`y,x" + "𝖽" X ball + using pmetric_space_pmetric_space1_valid by simp_all text‹The metric_space1› locale (context) specializes the the pmetric_space1› context by adding the assumption of identity of indiscernibles. › locale metric_space1 = pmetric_space1 + - assumes ident_indisc: "xX. yY. 𝖽`x,y = 𝟬 x=y" + assumes ident_indisc: "xX. yY. 𝖽`x,y = 𝟬 x=y" text‹The propositions proven in the metric_space› context defined in Metric_Space_ZF› theory are valid in the metric_space1› context. › -lemma (in metric_space1) metric_space_metric_space1_valid: +lemma (in metric_space1) metric_space_metric_space1_valid: shows "metric_space(,Add,ROrd,𝖽,X)" unfolding metric_space_def metric_space_axioms_def - using pmetric_space_pmetric_space1_valid ident_indisc + using pmetric_space_pmetric_space1_valid ident_indisc by simp +text‹The metric_space1› context is a special case of the metric_space› context, + with fixed objects mapping the same as in the mapping between pmetric_space1› + and pmetric_space› above. › + +sublocale metric_space1 < metric_space + "" Add ROrd "𝟬" realadd lesseq sless nonnegative positiveset + "λx y. LeftDiv(,Add)`x,y" + "λx y. RightDiv(,Add)`y,x" + "𝖽" X ball +proof + from ident_indisc show "xX. yX. 𝖽 ` x, y = TheNeutralElement(, Add) x = y" + by simp +qed + subsection‹Metric spaces are Hausdorff as topological spaces› text‹The usual (real-valued) metric spaces are a special case of ordered loop valued metric spaces defined in the MetricSpace_ZF› theory, hence they are $T_2$ - as topological spaces. › + as topological spaces. Below we repeat the major theorems of MetricSpace_ZF› theory + specialized the standard setting of real valued metrics. › text‹Since in the pmetric_space1› context $\mathfrak{d}$ is a pseudometrics - the (p)metric topology as defined above is indeed a topology, - the set of open balls is the base of that topology and the carrier of the - topology is the underlying (p)metric space carrier $X$. › - -theorem (in pmetric_space1) rpmetric_is_top: - shows - "Metric_Topology {is a topology}" - "Open_Balls {is a base for} Metric_Topology" - " Metric_Topology = X" - unfolding Open_Balls_def Metric_Topology_def - using rord_down_directs pmetric_space_pmetric_space1_valid - pmetric_space.pmetric_is_top by simp_all - -text‹The topology generated by a metric is Hausdorff (i.e. $T_2$). › - -theorem (in metric_space1) rmetric_space_T2: shows "Metric_Topology {is T2}" - unfolding Open_Balls_def Metric_Topology_def - using rord_down_directs metric_space_metric_space1_valid - metric_space.metric_space_T2 by simp + the (real valued) metric topology indeed a topology. › + +theorem (in pmetric_space1) rpmetric_is_top: + shows "τ {is a topology}" + using rord_down_directs pmetric_is_top by simp + +text‹The collection of open disks (caled ball›s in the pmetric_space1› context + is a base for the (real valued) metric topology.› + +theorem (in pmetric_space1) rdisks_are_base: + shows "(cX. {ball(c,R). R+}) {is a base for} τ" + using rord_down_directs disks_are_base by simp + +text‹$X$ is the carrier of the (real valued) metric topology.› + +theorem (in pmetric_space1) rmetric_top_carrier: shows "τ = X" + using rord_down_directs metric_top_carrier by simp + +text‹The topology generated by a (real valued) metric is Hausdorff (i.e. $T_2$). › + +theorem (in metric_space1) rmetric_space_T2: shows "τ {is T2}" + using rord_down_directs metric_space_T2 by simp end diff --git a/docs/IsarMathLib/Order_ZF.html b/docs/IsarMathLib/Order_ZF.html index 4387bb8..02fd10a 100644 --- a/docs/IsarMathLib/Order_ZF.html +++ b/docs/IsarMathLib/Order_ZF.html @@ -317,32 +317,32 @@

Theory Order_ZF

split Order_ZF_2_L1› into two lemmas.› lemma Order_ZF_2_L1A: assumes "x Interval(r,a,b)" - shows " a,x r" " x,b r" + shows "a,x r" "x,b r" using assms Order_ZF_2_L1 by auto textOrder_ZF_2_L1›, implication from right to left.› -lemma Order_ZF_2_L1B: assumes " a,x r" " x,b r" +lemma Order_ZF_2_L1B: assumes "a,x r" "x,b r" shows "x Interval(r,a,b)" - using assms Order_ZF_2_L1 by simp + using assms Order_ZF_2_L1 by simp text‹If the relation is reflexive, the endpoints belong to the interval.› -lemma Order_ZF_2_L2: assumes "refl(X,r)" - and "aX" "bX" and " a,b r" +lemma Order_ZF_2_L2: assumes "refl(X,r)" + and "aX" "bX" and "a,b r" shows "a Interval(r,a,b)" "b Interval(r,a,b)" - using assms refl_def Order_ZF_2_L1 by auto + using assms refl_def Order_ZF_2_L1 by auto text‹Under the assumptions of Order_ZF_2_L2›, the interval is nonempty.› -lemma Order_ZF_2_L2A: assumes "refl(X,r)" - and "aX" "bX" and " a,b r" +lemma Order_ZF_2_L2A: assumes "refl(X,r)" + and "aX" "bX" and "a,b r" shows "Interval(r,a,b) 0" proof - - from assms have "a Interval(r,a,b)" + from assms have "a Interval(r,a,b)" using Order_ZF_2_L2 by simp then show "Interval(r,a,b) 0" by auto qed @@ -351,19 +351,19 @@

Theory Order_ZF

only need trasitivity for this to be true.›
lemma Order_ZF_2_L3: - assumes A1: "trans(r)" and A2:" a,br" " b,cr" " c,dr" + assumes A1: "trans(r)" and A2:"a,br" "b,cr" "c,dr" shows "Interval(r,b,c) Interval(r,a,d)" proof - fix x assume A3: "x Interval(r, b, c)" - note A1 - moreover from A2 A3 have " a,b r b,x r" using Order_ZF_2_L1A + fix x assume A3: "x Interval(r, b, c)" + note A1 + moreover from A2 A3 have "a,b r b,x r" using Order_ZF_2_L1A by simp - ultimately have T1: " a,x r" by (rule Fol1_L3) - note A1 - moreover from A2 A3 have " x,c r c,d r" using Order_ZF_2_L1A + ultimately have T1: "a,x r" by (rule Fol1_L3) + note A1 + moreover from A2 A3 have "x,c r c,d r" using Order_ZF_2_L1A by simp - ultimately have " x,d r" by (rule Fol1_L3) - with T1 show "x Interval(r,a,d)" using Order_ZF_2_L1B + ultimately have "x,d r" by (rule Fol1_L3) + with T1 show "x Interval(r,a,d)" using Order_ZF_2_L1B by simp qed @@ -371,33 +371,33 @@

Theory Order_ZF

endpoints consists only of that endpoint.›
lemma Order_ZF_2_L4: - assumes A1: "refl(X,r)" and A2: "antisym(r)" and A3: "aX" + assumes A1: "refl(X,r)" and A2: "antisym(r)" and A3: "aX" shows "Interval(r,a,a) = {a}" proof - from A1 A3 have " a,a r" using refl_def by simp - with A1 A3 show "{a} Interval(r,a,a)" using Order_ZF_2_L2 by simp - from A2 show "Interval(r,a,a) {a}" using Order_ZF_2_L1A Fol1_L4 + from A1 A3 have " a,a r" using refl_def by simp + with A1 A3 show "{a} Interval(r,a,a)" using Order_ZF_2_L2 by simp + from A2 show "Interval(r,a,a) {a}" using Order_ZF_2_L1A Fol1_L4 by fast qed text‹For transitive relations the endpoints have to be in the relation for the interval to be nonempty.› -lemma Order_ZF_2_L5: assumes A1: "trans(r)" and A2: " a,b r" +lemma Order_ZF_2_L5: assumes A1: "trans(r)" and A2: " a,b r" shows "Interval(r,a,b) = 0" proof - { assume "Interval(r,a,b)0" then obtain x where "x Interval(r,a,b)" by auto - with A1 A2 have False using Order_ZF_2_L1A Fol1_L3 by fast + with A1 A2 have False using Order_ZF_2_L1A Fol1_L3 by fast } thus ?thesis by auto qed text‹If a relation is defined on a set, then intervals are subsets of that set.› -lemma Order_ZF_2_L6: assumes A1: "r X×X" +lemma Order_ZF_2_L6: assumes A1: "r X×X" shows "Interval(r,a,b) X" - using assms Interval_def by auto + using assms Interval_def by auto subsection‹Bounded sets› @@ -405,59 +405,59 @@

Theory Order_ZF

text‹For reflexive relations singletons are bounded.› -lemma Order_ZF_3_L1: assumes "refl(X,r)" and "aX" +lemma Order_ZF_3_L1: assumes "refl(X,r)" and "aX" shows "IsBounded({a},r)" - using assms refl_def IsBoundedAbove_def IsBoundedBelow_def + using assms refl_def IsBoundedAbove_def IsBoundedBelow_def IsBounded_def by auto text‹Sets that are bounded above are contained in the domain of the relation.› -lemma Order_ZF_3_L1A: assumes "r X×X" +lemma Order_ZF_3_L1A: assumes "r X×X" and "IsBoundedAbove(A,r)" - shows "AX" using assms IsBoundedAbove_def by auto + shows "AX" using assms IsBoundedAbove_def by auto text‹Sets that are bounded below are contained in the domain of the relation.› -lemma Order_ZF_3_L1B: assumes "r X×X" +lemma Order_ZF_3_L1B: assumes "r X×X" and "IsBoundedBelow(A,r)" - shows "AX" using assms IsBoundedBelow_def by auto + shows "AX" using assms IsBoundedBelow_def by auto text‹For a total relation, the greater of two elements, as defined above, is indeed greater of any of the two.› -lemma Order_ZF_3_L2: assumes "r {is total on} X" +lemma Order_ZF_3_L2: assumes "r {is total on} X" and "xX" "yX" shows "x,GreaterOf(r,x,y) r" "y,GreaterOf(r,x,y) r" "SmallerOf(r,x,y),x r" "SmallerOf(r,x,y),y r" - using assms IsTotal_def Order_ZF_1_L1 GreaterOf_def SmallerOf_def + using assms IsTotal_def Order_ZF_1_L1 GreaterOf_def SmallerOf_def by auto text‹If $A$ is bounded above by $u$, $B$ is bounded above by $w$, then $A\cup B$ is bounded above by the greater of $u,w$.› lemma Order_ZF_3_L2B: - assumes A1: "r {is total on} X" and A2: "trans(r)" - and A3: "uX" "wX" - and A4: "xA. x,u r" "xB. x,w r" + assumes A1: "r {is total on} X" and A2: "trans(r)" + and A3: "uX" "wX" + and A4: "xA. x,u r" "xB. x,w r" shows "xAB. x,GreaterOf(r,u,w) r" proof let ?v = "GreaterOf(r,u,w)" - from A1 A3 have T1: " u,?v r" and T2: " w,?v r" + from A1 A3 have T1: " u,?v r" and T2: " w,?v r" using Order_ZF_3_L2 by auto - fix x assume A5: "xAB" show "x,?v r" + fix x assume A5: "xAB" show "x,?v r" proof - { assume "xA" - with A4 T1 have " x,u r u,?v r" by simp - with A2 have "x,?v r" by (rule Fol1_L3) } + with A4 T1 have " x,u r u,?v r" by simp + with A2 have "x,?v r" by (rule Fol1_L3) } moreover { assume "xA" - with A5 A4 T2 have " x,w r w,?v r" by simp - with A2 have "x,?v r" by (rule Fol1_L3) } + with A5 A4 T2 have " x,w r w,?v r" by simp + with A2 have "x,?v r" by (rule Fol1_L3) } ultimately show ?thesis by auto qed qed @@ -466,21 +466,21 @@

Theory Order_ZF

above is bounded above.›
lemma Order_ZF_3_L3: - assumes A1: "r {is total on} X" and A2: "trans(r)" - and A3: "IsBoundedAbove(A,r)" "IsBoundedAbove(B,r)" - and A4: "r X×X" + assumes A1: "r {is total on} X" and A2: "trans(r)" + and A3: "IsBoundedAbove(A,r)" "IsBoundedAbove(B,r)" + and A4: "r X×X" shows "IsBoundedAbove(AB,r)" proof - { assume "A=0 B=0" - with A3 have "IsBoundedAbove(AB,r)" by auto } + with A3 have "IsBoundedAbove(AB,r)" by auto } moreover { assume "¬ (A = 0 B = 0)" - then have T1: "A0" "B0" by auto - with A3 obtain u w where D1: "xA. x,u r" "xB. x,w r" + then have T1: "A0" "B0" by auto + with A3 obtain u w where D1: "xA. x,u r" "xB. x,w r" using IsBoundedAbove_def by auto let ?U = "GreaterOf(r,u,w)" - from T1 A4 D1 have "uX" "wX" by auto - with A1 A2 D1 have "xAB. x,?U r" + from T1 A4 D1 have "uX" "wX" by auto + with A1 A2 D1 have "xAB. x,?U r" using Order_ZF_3_L2B by blast then have "IsBoundedAbove(AB,r)" using IsBoundedAbove_def by auto } @@ -490,14 +490,14 @@

Theory Order_ZF

text‹For total and transitive relations if a set $A$ is bounded above then $A\cup \{a\}$ is bounded above.› -lemma Order_ZF_3_L4: - assumes A1: "r {is total on} X" and A2: "trans(r)" +lemma Order_ZF_3_L4: + assumes A1: "r {is total on} X" and A2: "trans(r)" and A3: "IsBoundedAbove(A,r)" and A4: "aX" and A5: "r X×X" shows "IsBoundedAbove(A{a},r)" proof - - from A1 have "refl(X,r)" + from A1 have "refl(X,r)" using total_is_refl by simp - with assms show ?thesis using + with assms show ?thesis using Order_ZF_3_L1 IsBounded_def Order_ZF_3_L3 by simp qed @@ -505,23 +505,23 @@

Theory Order_ZF

then $A\cup B$ is bounded below by the smaller of $u,w$.›
lemma Order_ZF_3_L5B: - assumes A1: "r {is total on} X" and A2: "trans(r)" - and A3: "lX" "mX" - and A4: "xA. l,x r" "xB. m,x r" + assumes A1: "r {is total on} X" and A2: "trans(r)" + and A3: "lX" "mX" + and A4: "xA. l,x r" "xB. m,x r" shows "xAB. SmallerOf(r,l,m),x r" proof let ?k = "SmallerOf(r,l,m)" - from A1 A3 have T1: " ?k,l r" and T2: " ?k,m r" + from A1 A3 have T1: " ?k,l r" and T2: " ?k,m r" using Order_ZF_3_L2 by auto - fix x assume A5: "xAB" show "?k,x r" + fix x assume A5: "xAB" show "?k,x r" proof - { assume "xA" - with A4 T1 have " ?k,l r l,x r" by simp - with A2 have "?k,x r" by (rule Fol1_L3) } + with A4 T1 have " ?k,l r l,x r" by simp + with A2 have "?k,x r" by (rule Fol1_L3) } moreover { assume "xA" - with A5 A4 T2 have " ?k,m r m,x r" by simp - with A2 have "?k,x r" by (rule Fol1_L3) } + with A5 A4 T2 have " ?k,m r m,x r" by simp + with A2 have "?k,x r" by (rule Fol1_L3) } ultimately show ?thesis by auto qed qed @@ -530,21 +530,21 @@

Theory Order_ZF

below is bounded below.›
lemma Order_ZF_3_L6: - assumes A1: "r {is total on} X" and A2: "trans(r)" - and A3: "IsBoundedBelow(A,r)" "IsBoundedBelow(B,r)" - and A4: "r X×X" + assumes A1: "r {is total on} X" and A2: "trans(r)" + and A3: "IsBoundedBelow(A,r)" "IsBoundedBelow(B,r)" + and A4: "r X×X" shows "IsBoundedBelow(AB,r)" proof - { assume "A=0 B=0" - with A3 have ?thesis by auto } + with A3 have ?thesis by auto } moreover { assume "¬ (A = 0 B = 0)" - then have T1: "A0" "B0" by auto - with A3 obtain l m where D1: "xA. l,x r" "xB. m,x r" + then have T1: "A0" "B0" by auto + with A3 obtain l m where D1: "xA. l,x r" "xB. m,x r" using IsBoundedBelow_def by auto let ?L = "SmallerOf(r,l,m)" - from T1 A4 D1 have T1: "lX" "mX" by auto - with A1 A2 D1 have "xAB. ?L,x r" + from T1 A4 D1 have T1: "lX" "mX" by auto + with A1 A2 D1 have "xAB. ?L,x r" using Order_ZF_3_L5B by blast then have "IsBoundedBelow(AB,r)" using IsBoundedBelow_def by auto } @@ -554,43 +554,43 @@

Theory Order_ZF

text‹For total and transitive relations if a set $A$ is bounded below then $A\cup \{a\}$ is bounded below.› -lemma Order_ZF_3_L7: - assumes A1: "r {is total on} X" and A2: "trans(r)" +lemma Order_ZF_3_L7: + assumes A1: "r {is total on} X" and A2: "trans(r)" and A3: "IsBoundedBelow(A,r)" and A4: "aX" and A5: "r X×X" shows "IsBoundedBelow(A{a},r)" proof - - from A1 have "refl(X,r)" + from A1 have "refl(X,r)" using total_is_refl by simp - with assms show ?thesis using + with assms show ?thesis using Order_ZF_3_L1 IsBounded_def Order_ZF_3_L6 by simp qed text‹For total and transitive relations unions of two bounded sets are bounded.› -theorem Order_ZF_3_T1: +theorem Order_ZF_3_T1: assumes "r {is total on} X" and "trans(r)" and "IsBounded(A,r)" "IsBounded(B,r)" and "r X×X" shows "IsBounded(AB,r)" - using assms Order_ZF_3_L3 Order_ZF_3_L6 Order_ZF_3_L7 IsBounded_def + using assms Order_ZF_3_L3 Order_ZF_3_L6 Order_ZF_3_L7 IsBounded_def by simp text‹For total and transitive relations if a set $A$ is bounded then $A\cup \{a\}$ is bounded.› -lemma Order_ZF_3_L8: +lemma Order_ZF_3_L8: assumes "r {is total on} X" and "trans(r)" and "IsBounded(A,r)" and "aX" and "r X×X" shows "IsBounded(A{a},r)" - using assms total_is_refl Order_ZF_3_L1 Order_ZF_3_T1 by blast + using assms total_is_refl Order_ZF_3_L1 Order_ZF_3_T1 by blast text‹A sufficient condition for a set to be bounded below.› -lemma Order_ZF_3_L9: assumes A1: "aA. l,a r" +lemma Order_ZF_3_L9: assumes A1: "aA. l,a r" shows "IsBoundedBelow(A,r)" proof - - from A1 have "l. xA. l,x r" + from A1 have "l. xA. l,x r" by auto then show "IsBoundedBelow(A,r)" using IsBoundedBelow_def by simp @@ -598,10 +598,10 @@

Theory Order_ZF

text‹A sufficient condition for a set to be bounded above.› -lemma Order_ZF_3_L10: assumes A1: "aA. a,u r" +lemma Order_ZF_3_L10: assumes A1: "aA. a,u r" shows "IsBoundedAbove(A,r)" proof - - from A1 have "u. xA. x,u r" + from A1 have "u. xA. x,u r" by auto then show "IsBoundedAbove(A,r)" using IsBoundedAbove_def by simp @@ -631,17 +631,17 @@

Theory Order_ZF

text‹A subset of a set that is bounded below is bounded below.› -lemma Order_ZF_3_L12: assumes A1: "IsBoundedBelow(A,r)" and A2: "BA" +lemma Order_ZF_3_L12: assumes A1: "IsBoundedBelow(A,r)" and A2: "BA" shows "IsBoundedBelow(B,r)" proof - { assume "A = 0" - with assms have "IsBoundedBelow(B,r)" + with assms have "IsBoundedBelow(B,r)" using IsBoundedBelow_def by auto } moreover { assume "A 0" - with A1 have "l. xA. l,x r" + with A1 have "l. xA. l,x r" using IsBoundedBelow_def by simp - with A2 have "l.xB. l,x r" by auto + with A2 have "l.xB. l,x r" by auto then have "IsBoundedBelow(B,r)" using IsBoundedBelow_def by auto } ultimately show "IsBoundedBelow(B,r)" by auto @@ -649,17 +649,17 @@

Theory Order_ZF

text‹A subset of a set that is bounded above is bounded above.› -lemma Order_ZF_3_L13: assumes A1: "IsBoundedAbove(A,r)" and A2: "BA" +lemma Order_ZF_3_L13: assumes A1: "IsBoundedAbove(A,r)" and A2: "BA" shows "IsBoundedAbove(B,r)" proof - { assume "A = 0" - with assms have "IsBoundedAbove(B,r)" + with assms have "IsBoundedAbove(B,r)" using IsBoundedAbove_def by auto } moreover { assume "A 0" - with A1 have "u. xA. x,u r" + with A1 have "u. xA. x,u r" using IsBoundedAbove_def by simp - with A2 have "u.xB. x,u r" by auto + with A2 have "u.xB. x,u r" by auto then have "IsBoundedAbove(B,r)" using IsBoundedAbove_def by auto } ultimately show "IsBoundedAbove(B,r)" by auto @@ -672,21 +672,21 @@

Theory Order_ZF

lemma Order_ZF_3_L14: assumes A1: "r {is total on} X" - and A2: "trans(r)" and A3: "antisym(r)" - and A4: "r X×X" and A5: "X0" - and A6: "xX. aA. xa x,a r" + and A2: "trans(r)" and A3: "antisym(r)" + and A4: "r X×X" and A5: "X0" + and A6: "xX. aA. xa x,a r" shows "¬IsBoundedAbove(A,r)" proof - - { from A5 A6 have I: "A0" by auto + { from A5 A6 have I: "A0" by auto moreover assume "IsBoundedAbove(A,r)" - ultimately obtain u where II: "xA. x,u r" + ultimately obtain u where II: "xA. x,u r" using IsBounded_def IsBoundedAbove_def by auto - with A4 I have "uX" by auto - with A6 obtain b where "bA" and III: "ub" and "u,b r" + with A4 I have "uX" by auto + with A6 obtain b where "bA" and III: "ub" and "u,b r" by auto - with II have "b,u r" "u,b r" by auto - with A3 have "b=u" by (rule Fol1_L4) - with III have False by simp + with II have "b,u r" "u,b r" by auto + with A3 have "b=u" by (rule Fol1_L4) + with III have False by simp } thus "¬IsBoundedAbove(A,r)" by auto qed @@ -699,7 +699,7 @@

Theory Order_ZF

text‹If $A$ is bounded below, then the set of elements in a set $A$ that are nongreater than a given element is bounded.› -lemma Order_ZF_3_L16: assumes A1: "IsBoundedBelow(A,r)" +lemma Order_ZF_3_L16: assumes A1: "IsBoundedBelow(A,r)" shows "IsBounded({xA. x,a r},r)" proof - { assume "A=0" @@ -708,7 +708,7 @@

Theory Order_ZF

by auto } moreover { assume "A0" - with A1 obtain l where I: "xA. l,x r" + with A1 obtain l where I: "xA. l,x r" using IsBoundedBelow_def by auto then have "y{xA. x,a r}. l,y r" by simp then have "IsBoundedBelow({xA. x,a r},r)" diff --git a/docs/IsarMathLib/OrderedLoop_ZF.html b/docs/IsarMathLib/OrderedLoop_ZF.html index a6d46e4..053a9c7 100644 --- a/docs/IsarMathLib/OrderedLoop_ZF.html +++ b/docs/IsarMathLib/OrderedLoop_ZF.html @@ -118,45 +118,58 @@

Theory OrderedLoop_ZF

sublocale loop1 < loop0 L A looper using ordLoopAssum loop_loop0_valid unfolding IsAnOrdLoop_def by auto +text‹The notation $-x+y$ and $x-y$ denotes left and right division, resp. + These two operations are closed in a loop, see lemma lrdiv_binop› in the Quasigroup_ZF› theory. + The next lemma reiterates that fact using the notation of the loop1› context. › + +lemma (in loop1) left_right_sub_closed: assumes "xL" "yL" + shows "(\<rm>x\<ad>y) L" and "x\<rs>y L" +proof - + from qgroupassum have "LeftDiv(L,A):L×L L" and "RightDiv(L,A):L×L L" + using lrdiv_binop by simp_all + with assms show "(\<rm>x\<ad>y) L" and "x\<rs>y L" + using apply_funtype by simp_all +qed + text‹In this context $x \leq y$ implies that both $x$ and $y$ belong to $L$.› -lemma (in loop1) lsq_members: assumes "x\<lsq>y" shows "xL" and "yL" - using ordLoopAssum assms IsAnOrdLoop_def by auto +lemma (in loop1) lsq_members: assumes "x\<lsq>y" shows "xL" and "yL" + using ordLoopAssum assms IsAnOrdLoop_def by auto text‹In this context $x < y$ implies that both $x$ and $y$ belong to $L$.› -lemma (in loop1) less_members: assumes "x\<ls>y" shows "xL" and "yL" - using ordLoopAssum assms IsAnOrdLoop_def by auto +lemma (in loop1) less_members: assumes "x\<ls>y" shows "xL" and "yL" + using ordLoopAssum assms IsAnOrdLoop_def by auto text‹ In an ordered loop the order is translation invariant. › -lemma (in loop1) ord_trans_inv: assumes "x\<lsq>y" "zL" +lemma (in loop1) ord_trans_inv: assumes "x\<lsq>y" "zL" shows "x\<ra>z \<lsq> y\<ra>z" and "z\<ra>x \<lsq> z\<ra>y" proof - - from ordLoopAssum assms have + from ordLoopAssum assms have "(x,y r A` x,z,A`y,z r) (x,y r A`z,x,A`z,y r )" - using lsq_members unfolding IsAnOrdLoop_def by blast - with assms(1) show "x\<ra>z \<lsq> y\<ra>z" and "z\<ra>x \<lsq> z\<ra>y" by auto + using lsq_members unfolding IsAnOrdLoop_def by blast + with assms(1) show "x\<ra>z \<lsq> y\<ra>z" and "z\<ra>x \<lsq> z\<ra>y" by auto qed text‹ In an ordered loop the strict order is translation invariant. › -lemma (in loop1) strict_ord_trans_inv: assumes "x\<ls>y" "zL" +lemma (in loop1) strict_ord_trans_inv: assumes "x\<ls>y" "zL" shows "x\<ra>z \<ls> y\<ra>z" and "z\<ra>x \<ls> z\<ra>y" proof - - from assms have "x\<ra>z \<lsq> y\<ra>z" and "z\<ra>x \<lsq> z\<ra>y" - using ord_trans_inv by auto + from assms have "x\<ra>z \<lsq> y\<ra>z" and "z\<ra>x \<lsq> z\<ra>y" + using ord_trans_inv by auto moreover have "x\<ra>z y\<ra>z" and "z\<ra>x z\<ra>y" proof - { assume "x\<ra>z = y\<ra>z" - with assms have "x=y" using less_members qg_cancel_right by blast - with assms(1) have False by simp + with assms have "x=y" using less_members qg_cancel_right by blast + with assms(1) have False by simp } thus "x\<ra>z y\<ra>z" by auto { assume "z\<ra>x = z\<ra>y" - with assms have "x=y" using less_members qg_cancel_left by blast - with assms(1) have False by simp + with assms have "x=y" using less_members qg_cancel_left by blast + with assms(1) have False by simp } thus "z\<ra>x z\<ra>y" by auto qed ultimately show "x\<ra>z \<ls> y\<ra>z" and "z\<ra>x \<ls> z\<ra>y" @@ -165,40 +178,40 @@

Theory OrderedLoop_ZF

text‹We can cancel an element from both sides of an inequality on the right side. › -lemma (in loop1) ineq_cancel_right: assumes "xL" "yL" "zL" and "x\<ra>z \<lsq> y\<ra>z" +lemma (in loop1) ineq_cancel_right: assumes "xL" "yL" "zL" and "x\<ra>z \<lsq> y\<ra>z" shows "x\<lsq>y" proof - - from ordLoopAssum assms(1,2,3) have "x,y r A` x,z,A`y,z r" + from ordLoopAssum assms(1,2,3) have "x,y r A` x,z,A`y,z r" unfolding IsAnOrdLoop_def by blast - with assms(4) show "x\<lsq>y" by simp + with assms(4) show "x\<lsq>y" by simp qed text‹We can cancel an element from both sides of a strict inequality on the right side. › -lemma (in loop1) strict_ineq_cancel_right: assumes "xL" "yL" "zL" and "x\<ra>z \<ls> y\<ra>z" +lemma (in loop1) strict_ineq_cancel_right: assumes "xL" "yL" "zL" and "x\<ra>z \<ls> y\<ra>z" shows "x\<ls>y" - using assms ineq_cancel_right by auto + using assms ineq_cancel_right by auto text‹We can cancel an element from both sides of an inequality on the left side. › -lemma (in loop1) ineq_cancel_left: assumes "xL" "yL" "zL" and "z\<ra>x \<lsq> z\<ra>y" +lemma (in loop1) ineq_cancel_left: assumes "xL" "yL" "zL" and "z\<ra>x \<lsq> z\<ra>y" shows "x\<lsq>y" proof - - from ordLoopAssum assms(1,2,3) have "x,y r A`z,x,A`z,y r" + from ordLoopAssum assms(1,2,3) have "x,y r A`z,x,A`z,y r" unfolding IsAnOrdLoop_def by blast - with assms(4) show "x\<lsq>y" by simp + with assms(4) show "x\<lsq>y" by simp qed text‹We can cancel an element from both sides of a strict inequality on the left side. › -lemma (in loop1) strict_ineq_cancel_left: +lemma (in loop1) strict_ineq_cancel_left: assumes "xL" "yL" "zL" and "z\<ra>x \<ls> z\<ra>y" shows "x\<ls>y" - using assms ineq_cancel_left by auto + using assms ineq_cancel_left by auto text‹The definition of the nonnegative set in the notation used in the loop1› locale: › -lemma (in loop1) nonneg_definition: +lemma (in loop1) nonneg_definition: shows "x L+ 𝟬 \<lsq> x" using ordLoopAssum IsAnOrdLoop_def Nonnegative_def by auto text‹The nonnegative set is contained in the loop.› @@ -219,169 +232,181 @@

Theory OrderedLoop_ZF

text‹Another form of the definition of the positive set in the notation used in the loop1› locale: › -lemma (in loop1) posset_definition1: +lemma (in loop1) posset_definition1: shows "x L+ 𝟬\<ls>x" using ordLoopAssum IsAnOrdLoop_def PositiveSet_def by auto text‹ The order in an ordered loop is antisymmeric. › -lemma (in loop1) loop_ord_antisym: assumes "x\<lsq>y" and "y\<lsq>x" +lemma (in loop1) loop_ord_antisym: assumes "x\<lsq>y" and "y\<lsq>x" shows "x=y" proof - - from ordLoopAssum assms have "antisym(r)" "x,y r" "y,x r" + from ordLoopAssum assms have "antisym(r)" "x,y r" "y,x r" unfolding IsAnOrdLoop_def IsPartOrder_def by auto then show "x=y" by (rule Fol1_L4) qed text‹ The loop order is transitive. › -lemma (in loop1) loop_ord_trans: assumes "x\<lsq>y" and "y\<lsq>z" shows "x\<lsq>z" +lemma (in loop1) loop_ord_trans: assumes "x\<lsq>y" and "y\<lsq>z" shows "x\<lsq>z" proof - - from ordLoopAssum assms have "trans(r)" and "x,y r y,z r" + from ordLoopAssum assms have "trans(r)" and "x,y r y,z r" unfolding IsAnOrdLoop_def IsPartOrder_def by auto then have "x,z r" by (rule Fol1_L3) thus ?thesis by simp qed text‹ The loop order is reflexive. › -lemma (in loop1) loop_ord_refl: assumes "xL" shows "x\<lsq>x" - using assms ordLoopAssum unfolding IsAnOrdLoop_def IsPartOrder_def refl_def + +lemma (in loop1) loop_ord_refl: assumes "xL" shows "x\<lsq>x" + using assms ordLoopAssum unfolding IsAnOrdLoop_def IsPartOrder_def refl_def + by simp + +text‹The neutral element is nonnegative.› + +lemma (in loop1) loop_zero_nonneg: shows "𝟬L+" + using neut_props_loop(1) loop_ord_refl nonneg_definition by simp text‹ A form of mixed transitivity for the strict order: › -lemma (in loop1) loop_strict_ord_trans: assumes "x\<lsq>y" and "y\<ls>z" +lemma (in loop1) loop_strict_ord_trans: assumes "x\<lsq>y" and "y\<ls>z" shows "x\<ls>z" proof - - from assms have "x\<lsq>y" and "y\<lsq>z" by auto - then have "x\<lsq>z" by (rule loop_ord_trans) - with assms show "x\<ls>z" using loop_ord_antisym by auto + from assms have "x\<lsq>y" and "y\<lsq>z" by auto + then have "x\<lsq>z" by (rule loop_ord_trans) + with assms show "x\<ls>z" using loop_ord_antisym by auto qed text‹ Another form of mixed transitivity for the strict order: › -lemma (in loop1) loop_strict_ord_trans1: assumes "x\<ls>y" and "y\<lsq>z" +lemma (in loop1) loop_strict_ord_trans1: assumes "x\<ls>y" and "y\<lsq>z" shows "x\<ls>z" proof - - from assms have "x\<lsq>y" and "y\<lsq>z" by auto - then have "x\<lsq>z" by (rule loop_ord_trans) - with assms show "x\<ls>z" using loop_ord_antisym by auto + from assms have "x\<lsq>y" and "y\<lsq>z" by auto + then have "x\<lsq>z" by (rule loop_ord_trans) + with assms show "x\<ls>z" using loop_ord_antisym by auto qed text‹ Yet another form of mixed transitivity for the strict order: › -lemma (in loop1) loop_strict_ord_trans2: assumes "x\<ls>y" and "y\<ls>z" +lemma (in loop1) loop_strict_ord_trans2: assumes "x\<ls>y" and "y\<ls>z" shows "x\<ls>z" proof - - from assms have "x\<lsq>y" and "y\<lsq>z" by auto - then have "x\<lsq>z" by (rule loop_ord_trans) - with assms show "x\<ls>z" using loop_ord_antisym by auto + from assms have "x\<lsq>y" and "y\<lsq>z" by auto + then have "x\<lsq>z" by (rule loop_ord_trans) + with assms show "x\<ls>z" using loop_ord_antisym by auto qed text‹ We can move an element to the other side of an inequality. Well, not exactly, but our notation creates an illusion to that effect. › -lemma (in loop1) lsq_other_side: assumes "x\<lsq>y" +lemma (in loop1) lsq_other_side: assumes "x\<lsq>y" shows "𝟬 \<lsq> \<rm>x\<ad>y" "(\<rm>x\<ad>y) L+" "𝟬 \<lsq> y\<rs>x" "(y\<rs>x) L+" proof - - from assms have "xL" "yL" "𝟬L" "(\<rm>x\<ad>y) L" "(y\<rs>x) L" - using lsq_members neut_props_loop(1) lrdiv_props(2,5) by auto + from assms have "xL" "yL" "𝟬L" "(\<rm>x\<ad>y) L" "(y\<rs>x) L" + using lsq_members neut_props_loop(1) lrdiv_props(2,5) by auto then have "x = x\<ra>𝟬" and "y = x\<ra>(\<rm>x\<ad>y)" using neut_props_loop(2) lrdiv_props(6) by auto - with assms have "x\<ra>𝟬 \<lsq> x\<ra>(\<rm>x\<ad>y)" by simp - with xL 𝟬L (\<rm>x\<ad>y) L show "𝟬 \<lsq> \<rm>x\<ad>y" using ineq_cancel_left + with assms have "x\<ra>𝟬 \<lsq> x\<ra>(\<rm>x\<ad>y)" by simp + with xL 𝟬L (\<rm>x\<ad>y) L show "𝟬 \<lsq> \<rm>x\<ad>y" using ineq_cancel_left by simp - then show "(\<rm>x\<ad>y) L+" using nonneg_definition by simp + then show "(\<rm>x\<ad>y) L+" using nonneg_definition by simp from xL yL have "x = 𝟬\<ra>x" and "y = (y\<rs>x)\<ra>x" using neut_props_loop(2) lrdiv_props(3) by auto - with assms have "𝟬\<ra>x \<lsq> (y\<rs>x)\<ra>x" by simp - with xL 𝟬L (y\<rs>x) L show "𝟬 \<lsq> y\<rs>x" using ineq_cancel_right + with assms have "𝟬\<ra>x \<lsq> (y\<rs>x)\<ra>x" by simp + with xL 𝟬L (y\<rs>x) L show "𝟬 \<lsq> y\<rs>x" using ineq_cancel_right by simp - then show "(y\<rs>x) L+" using nonneg_definition by simp + then show "(y\<rs>x) L+" using nonneg_definition by simp qed text‹ We can move an element to the other side of a strict inequality. › -lemma (in loop1) ls_other_side: assumes "x\<ls>y" +lemma (in loop1) ls_other_side: assumes "x\<ls>y" shows "𝟬 \<ls> \<rm>x\<ad>y" "(\<rm>x\<ad>y) L+" "𝟬 \<ls> y\<rs>x" "(y\<rs>x) L+" proof - - from assms have "xL" "yL" "𝟬L" "(\<rm>x\<ad>y) L" "(y\<rs>x) L" - using lsq_members neut_props_loop(1) lrdiv_props(2,5) by auto + from assms have "xL" "yL" "𝟬L" "(\<rm>x\<ad>y) L" "(y\<rs>x) L" + using lsq_members neut_props_loop(1) lrdiv_props(2,5) by auto then have "x = x\<ra>𝟬" and "y = x\<ra>(\<rm>x\<ad>y)" using neut_props_loop(2) lrdiv_props(6) by auto - with assms have "x\<ra>𝟬 \<ls> x\<ra>(\<rm>x\<ad>y)" by simp - with xL 𝟬L (\<rm>x\<ad>y) L show "𝟬 \<ls> \<rm>x\<ad>y" using strict_ineq_cancel_left + with assms have "x\<ra>𝟬 \<ls> x\<ra>(\<rm>x\<ad>y)" by simp + with xL 𝟬L (\<rm>x\<ad>y) L show "𝟬 \<ls> \<rm>x\<ad>y" using strict_ineq_cancel_left by simp - then show "(\<rm>x\<ad>y) L+" using posset_definition1 by simp + then show "(\<rm>x\<ad>y) L+" using posset_definition1 by simp from xL yL have "x = 𝟬\<ra>x" and "y = (y\<rs>x)\<ra>x" using neut_props_loop(2) lrdiv_props(3) by auto - with assms have "𝟬\<ra>x \<ls> (y\<rs>x)\<ra>x" by simp - with xL 𝟬L (y\<rs>x) L show "𝟬 \<ls> y\<rs>x" using strict_ineq_cancel_right + with assms have "𝟬\<ra>x \<ls> (y\<rs>x)\<ra>x" by simp + with xL 𝟬L (y\<rs>x) L show "𝟬 \<ls> y\<rs>x" using strict_ineq_cancel_right by simp - then show "(y\<rs>x) L+" using posset_definition1 by simp + then show "(y\<rs>x) L+" using posset_definition1 by simp qed text‹We can add sides of inequalities.› -lemma (in loop1) add_ineq: assumes "x\<lsq>y" "z\<lsq>t" +lemma (in loop1) add_ineq: assumes "x\<lsq>y" "z\<lsq>t" shows "x\<ra>z \<lsq> y\<ra>t" proof - - from assms have "x\<ra>z \<lsq> y\<ra>z" - using lsq_members(1) ord_trans_inv(1) by simp - with assms show ?thesis using lsq_members(2) ord_trans_inv(2) loop_ord_trans + from assms have "x\<ra>z \<lsq> y\<ra>z" + using lsq_members(1) ord_trans_inv(1) by simp + with assms show ?thesis using lsq_members(2) ord_trans_inv(2) loop_ord_trans by simp qed text‹We can add sides of strict inequalities. The proof uses a lemma that relies on the antisymmetry of the order relation.› -lemma (in loop1) add_ineq_strict: assumes "x\<ls>y" "z\<ls>t" +lemma (in loop1) add_ineq_strict: assumes "x\<ls>y" "z\<ls>t" shows "x\<ra>z \<ls> y\<ra>t" proof - - from assms have "x\<ra>z \<ls> y\<ra>z" - using less_members(1) strict_ord_trans_inv(1) by auto - moreover from assms have "y\<ra>z \<ls> y\<ra>t" - using less_members(2) strict_ord_trans_inv(2) by auto - ultimately show ?thesis by (rule loop_strict_ord_trans2) + from assms have "x\<ra>z \<ls> y\<ra>z" + using less_members(1) strict_ord_trans_inv(1) by auto + moreover from assms have "y\<ra>z \<ls> y\<ra>t" + using less_members(2) strict_ord_trans_inv(2) by auto + ultimately show ?thesis by (rule loop_strict_ord_trans2) qed text‹We can add sides of inequalities one of which is strict. › -lemma (in loop1) add_ineq_strict1: assumes "x\<lsq>y" "z\<ls>t" +lemma (in loop1) add_ineq_strict1: assumes "x\<lsq>y" "z\<ls>t" shows "x\<ra>z \<ls> y\<ra>t" and "z\<ra>x \<ls> t\<ra>y" proof - - from assms have "x\<ra>z \<lsq> y\<ra>z" - using less_members(1) ord_trans_inv(1) by auto - with assms show "x\<ra>z \<ls> y\<ra>t" - using lsq_members(2) strict_ord_trans_inv(2) loop_strict_ord_trans + from assms have "x\<ra>z \<lsq> y\<ra>z" + using less_members(1) ord_trans_inv(1) by auto + with assms show "x\<ra>z \<ls> y\<ra>t" + using lsq_members(2) strict_ord_trans_inv(2) loop_strict_ord_trans by blast - from assms have "z\<ra>x \<ls> t\<ra>x" - using lsq_members(1) strict_ord_trans_inv(1) by simp - with assms show "z\<ra>x \<ls> t\<ra>y" - using less_members(2) ord_trans_inv(2) loop_strict_ord_trans1 + from assms have "z\<ra>x \<ls> t\<ra>x" + using lsq_members(1) strict_ord_trans_inv(1) by simp + with assms show "z\<ra>x \<ls> t\<ra>y" + using less_members(2) ord_trans_inv(2) loop_strict_ord_trans1 by blast qed -text‹Subtracting a positive element decreases the value. › +text‹Subtracting a positive element decreases the value, while adding a positive element + increases the value. › -lemma (in loop1) subtract_pos: assumes "xL" "𝟬\<ls>y" - shows "x\<rs>y \<ls> x" and "(\<rm>y\<ad>x) \<ls> x" +lemma (in loop1) add_subtract_pos: assumes "xL" "𝟬\<ls>y" + shows + "x\<rs>y \<ls> x" "(\<rm>y\<ad>x) \<ls> x" "x \<ls> x\<ra>y" "x \<ls> y\<ra>x" proof - - from assms(2) have "yL" using less_members(2) by simp - from assms(1) have "x\<lsq>x" + from assms(2) have "yL" using less_members(2) by simp + from assms(1) have "x\<lsq>x" using ordLoopAssum unfolding IsAnOrdLoop_def IsPartOrder_def refl_def by simp - with assms(2) have "x\<ra>𝟬 \<ls> x\<ra>y" - using add_ineq_strict1(1) by simp - with assms yL show "x\<rs>y \<ls> x" - using neut_props_loop(2) lrdiv_props(3) lrdiv_props(2) strict_ineq_cancel_right + with assms(2) have "x\<ra>𝟬 \<ls> x\<ra>y" + using add_ineq_strict1(1) by simp + with assms yL show "x\<rs>y \<ls> x" + using neut_props_loop(2) lrdiv_props(3) lrdiv_props(2) strict_ineq_cancel_right by simp - from assms(2) x\<lsq>x have "𝟬\<ra>x \<ls> y\<ra>x" - using add_ineq_strict1(2) by simp - with assms yL show "(\<rm>y\<ad>x) \<ls> x" - using neut_props_loop(2) lrdiv_props(6) lrdiv_props(5) strict_ineq_cancel_left + from assms(2) x\<lsq>x have "𝟬\<ra>x \<ls> y\<ra>x" + using add_ineq_strict1(2) by simp + with assms yL show "(\<rm>y\<ad>x) \<ls> x" + using neut_props_loop(2) lrdiv_props(6) lrdiv_props(5) strict_ineq_cancel_left by simp -qed + from assms(1) x\<ra>𝟬 \<ls> x\<ra>y 𝟬\<ra>x \<ls> y\<ra>x + show "x \<ls> x\<ra>y" "x \<ls> y\<ra>x" + using neut_props_loop(2) by simp_all +qed end diff --git a/docs/IsarMathLib/Real_ZF_2.html b/docs/IsarMathLib/Real_ZF_2.html index 339bf4a..fe09386 100644 --- a/docs/IsarMathLib/Real_ZF_2.html +++ b/docs/IsarMathLib/Real_ZF_2.html @@ -270,19 +270,25 @@

Theory Real_ZF_2

using pos_is_lattice(3) pos_non_empty meet_down_directs down_dir_mono unfolding IsAlattice_def by blast -text‹ We define the topology on reals as one consisting of the unions of open disks. › +text‹ We define the topology on reals as the metric topology + coming from the dist› metric (i.e. consisting of the unions of open disks). › -definition (in reals) RealTopology ("τ") - where "τ {A. A Pow(c.{disk(c,r). r +})}" +definition (in reals) RealTopology ("τ") + where "τ MetricTopology(,,Add,ROrd,dist)" + +text‹A more explicit definition of the real topology in notation used in the reals› context. › + +lemma (in reals) real_toplology_def_alt: + shows "τ = {A. A Pow(c.{disk(c,r). r +})}" + unfolding MetricTopology_def RealTopology_def by simp text‹Real numbers form a Hausdorff topological space with topology generated by open disks. › -theorem (in reals) reals_is_top: - shows "τ {is a topology}" "τ = " "τ {is T2}" - using rord_down_directs metric_space_valid pmetric_space_valid - pmetric_space.pmetric_is_top metric_space.metric_space_T2 - unfolding RealTopology_def - by simp_all +theorem (in reals) reals_is_top: + shows "τ {is a topology}" "τ = " "τ {is T2}" + using rord_down_directs metric_space_valid pmetric_space_valid pmetric_space.pmetric_is_top + pmetric_space.metric_top_carrier metric_space.metric_space_T2 + unfolding RealTopology_def by simp_all end diff --git a/docs/IsarMathLib/Topology_ZF_examples.html b/docs/IsarMathLib/Topology_ZF_examples.html index b61aef8..6d9a993 100644 --- a/docs/IsarMathLib/Topology_ZF_examples.html +++ b/docs/IsarMathLib/Topology_ZF_examples.html @@ -238,57 +238,55 @@

Theory Topology_ZF_examples

ultimately show ?thesis by auto qed -text‹ $X$ is a closed set that contains $A$. -This lemma is necessary because we cannot -use the lemmas proven in the topology0› context since - T≠0"} › is too weak for - CoCardinal(X,T)› to be a topology.› +text‹ $X$ is a closed set that contains $A$. + This lemma is necessary because we cannot use the lemmas proven in the + topology0› context since $T\neq 0$ is too weak for CoCardinal(X,T)› + to be a topology.› -lemma X_closedcov_cocardinal: +lemma X_closedcov_cocardinal: assumes "T0" "AX" - shows "XClosedCovers(A,CoCardinal(X,T))" using ClosedCovers_def - using union_cocardinal closed_sets_cocardinal assms by auto + shows "X ClosedCovers(A,CoCardinal(X,T))" using ClosedCovers_def + using union_cocardinal closed_sets_cocardinal assms by auto -text‹The closure of a set is itself if it is closed or X› if -it isn't closed.› +text‹The closure of a set is itself if it is closed or $X$ if it isn't closed.› -lemma closure_set_cocardinal: +lemma closure_set_cocardinal: assumes "T0""AX" shows "Closure(A,CoCardinal(X,T))=(if (A T) then A else X)" proof- { assume "A T" - with assms have "A {is closed in} CoCardinal(X,T)" using closed_sets_cocardinal by auto - with assms(2) have "A {D Pow(X). D {is closed in} CoCardinal(X,T) AD}" by auto - with assms(1) have S:"AClosedCovers(A,CoCardinal(X,T))" using ClosedCovers_def + with assms have "A {is closed in} CoCardinal(X,T)" using closed_sets_cocardinal by auto + with assms(2) have "A {D Pow(X). D {is closed in} CoCardinal(X,T) AD}" by auto + with assms(1) have S:"AClosedCovers(A,CoCardinal(X,T))" using ClosedCovers_def using union_cocardinal by auto - hence l1:"ClosedCovers(A,CoCardinal(X,T))A" by blast - from S have l2:"A ClosedCovers(A,CoCardinal(X,T))" + hence l1:"ClosedCovers(A,CoCardinal(X,T))A" by blast + from S have l2:"A ClosedCovers(A,CoCardinal(X,T))" unfolding ClosedCovers_def by auto - from l1 l2 have "Closure(A,CoCardinal(X,T))=A" using Closure_def + from l1 l2 have "Closure(A,CoCardinal(X,T))=A" using Closure_def by auto } moreover { - assume as:"¬ A T" + assume as:"¬ A T" { fix U assume "AU" - then have Q:"A U" using subset_imp_lepoll by auto + then have Q:"A U" using subset_imp_lepoll by auto { assume "U T" - with Q have "A T" using lesspoll_trans1 by auto - with as have "False" by auto + with Q have "A T" using lesspoll_trans1 by auto + with as have "False" by auto } hence "¬ U T" by auto - with assms(1) have "¬(U {is closed in} CoCardinal(X,T)) U=X" using closed_sets_cocardinal + with assms(1) have "¬(U {is closed in} CoCardinal(X,T)) U=X" using closed_sets_cocardinal by auto } - with assms(1) have "UPow(X). U{is closed in}CoCardinal(X,T) AUU=X" + with assms(1) have "UPow(X). U{is closed in}CoCardinal(X,T) AUU=X" by auto - with assms(1) have "ClosedCovers(A,CoCardinal(X,T)){X}" + with assms(1) have "ClosedCovers(A,CoCardinal(X,T)){X}" using union_cocardinal using ClosedCovers_def by auto - with assms have "ClosedCovers(A,CoCardinal(X,T))={X}" using X_closedcov_cocardinal + with assms have "ClosedCovers(A,CoCardinal(X,T))={X}" using X_closedcov_cocardinal by auto then have " Closure(A,CoCardinal(X,T)) = X " using Closure_def by auto } @@ -299,44 +297,44 @@

Theory Topology_ZF_examples

X› if not $A$ neither $X-A$ are closed and; if only one is closed, then the closed one is its boundary.› -lemma boundary_cocardinal: +lemma boundary_cocardinal: assumes "T0""AX" shows "Boundary(A,CoCardinal(X,T)) = (if A T then (if (X-A) T then 0 else A) else (if (X-A) T then X-A else X))" proof- - from assms(2) have "X-A X" by auto + from assms(2) have "X-A X" by auto { assume AS: "AT" "X-A T" - with assms X-A X have + with assms X-A X have "Closure(X-A,CoCardinal(X,T)) = X-A" and "Closure(A,CoCardinal(X,T)) = A" using closure_set_cocardinal by auto - with assms(1) have "Boundary(A,CoCardinal(X,T)) = 0" + with assms(1) have "Boundary(A,CoCardinal(X,T)) = 0" using Boundary_def union_cocardinal by auto } moreover { assume AS: "~(AT)" "X-A T" - with assms X-A X have + with assms X-A X have "Closure(X-A,CoCardinal(X,T)) = X-A" and "Closure(A,CoCardinal(X,T)) = X" using closure_set_cocardinal by auto - with assms(1) have "Boundary(A,CoCardinal(X,T))=X-A" using Boundary_def + with assms(1) have "Boundary(A,CoCardinal(X,T))=X-A" using Boundary_def union_cocardinal by auto } moreover { assume AS:"~(AT)" "~(X-A T)" - with assms X-A X have + with assms X-A X have "Closure(X-A,CoCardinal(X,T))=X" and "Closure(A,CoCardinal(X,T))=X" using closure_set_cocardinal by auto - with assms(1) have "Boundary(A,CoCardinal(X,T))=X" using Boundary_def union_cocardinal + with assms(1) have "Boundary(A,CoCardinal(X,T))=X" using Boundary_def union_cocardinal by auto } moreover { assume AS:"A T" "~(X-AT)" - with assms X-A X have + with assms X-A X have "Closure(X-A,CoCardinal(X,T))=X" and "Closure(A,CoCardinal(X,T)) = A" using closure_set_cocardinal by auto - with assms have "Boundary(A,CoCardinal(X,T))=A" using Boundary_def union_cocardinal + with assms have "Boundary(A,CoCardinal(X,T))=A" using Boundary_def union_cocardinal by auto } ultimately show ?thesis by auto @@ -345,7 +343,7 @@

Theory Topology_ZF_examples

text‹If the set is too small or the cardinal too large, then the topology is just the discrete topology.› -lemma discrete_cocardinal: +lemma discrete_cocardinal: assumes "XT" shows "CoCardinal(X,T) = Pow(X)" proof @@ -357,11 +355,11 @@

Theory Topology_ZF_examples

then show "CoCardinal(X,T) Pow(X)" by auto { fix U - assume A:"U Pow(X)" + assume A:"U Pow(X)" then have "X-U X" by auto then have "X-U X" using subset_imp_lepoll by auto - then have "X-U T" using lesspoll_trans1 assms by auto - with A have "UCoCardinal(X,T)" using CoCardinal_def + then have "X-U T" using lesspoll_trans1 assms by auto + with A have "UCoCardinal(X,T)" using CoCardinal_def by auto } then show "Pow(X) CoCardinal(X,T)" by auto @@ -396,13 +394,13 @@

Theory Topology_ZF_examples

{ fix M assume "M (CoCardinal(X,T) {restricted to} Y)" - then obtain A where A1:"A CoCardinal(X,T)" "M=Y A" using RestrictedTo_def by auto + then obtain A where A1:"A CoCardinal(X,T)" "M=Y A" using RestrictedTo_def by auto then have "M Pow(X Y)" using CoCardinal_def by auto moreover - from A1 have "(Y X)-M = (Y X)-A" using CoCardinal_def by auto + from A1 have "(Y X)-M = (Y X)-A" using CoCardinal_def by auto with (Y X)-M = (Y X)-A have "(Y X)-M X-A" by auto then have "(Y X)-M X-A" using subset_imp_lepoll by auto - with A1 have "(Y X)-M T M=0" using lesspoll_trans1 CoCardinal_def + with A1 have "(Y X)-M T M=0" using lesspoll_trans1 CoCardinal_def by auto ultimately have "M CoCardinal(YX, T)" using CoCardinal_def by auto @@ -411,7 +409,7 @@

Theory Topology_ZF_examples

{ fix M let ?A = "M (X-Y)" - assume A:"M CoCardinal(Y X,T)" + assume A:"M CoCardinal(Y X,T)" { assume "M=0" hence "M=0 Y" by auto @@ -420,19 +418,19 @@

Theory Topology_ZF_examples

} moreover { - assume AS:"M0" - from A AS have A1:"(MPow(Y X) (Y X)-MT)" using CoCardinal_def by auto + assume AS:"M0" + from A AS have A1:"(MPow(Y X) (Y X)-MT)" using CoCardinal_def by auto hence "?APow(X)" by blast moreover have "X-?A=(Y X)-M" by blast - with A1 have "X-?A T" by auto + with A1 have "X-?A T" by auto ultimately have "?ACoCardinal(X,T)" using CoCardinal_def by auto - then have AT:"Y ?ACoCardinal(X,T) {restricted to} Y" using RestrictedTo_def + then have AT:"Y ?ACoCardinal(X,T) {restricted to} Y" using RestrictedTo_def by auto have "Y ?A=Y M" by blast - also from A1 have "=M" by auto + also from A1 have "=M" by auto finally have "Y ?A=M" by simp - with AT have "MCoCardinal(X,T) {restricted to} Y" + with AT have "MCoCardinal(X,T) {restricted to} Y" by auto } ultimately have "MCoCardinal(X,T) {restricted to} Y" by auto @@ -458,20 +456,20 @@

Theory Topology_ZF_examples

{ fix M assume "M Pow(ExcludedSet(X,Q))" - then have A:"M{FPow(X). Q F=0} {X}" using ExcludedSet_def by auto + then have A:"M{FPow(X). Q F=0} {X}" using ExcludedSet_def by auto hence "MPow(X)" by auto moreover { - have B:"Q M={Q T. TM}" by auto + have B:"Q M={Q T. TM}" by auto { assume "XM" - with A have "M{FPow(X). Q F=0}" by auto - with B have "Q M=0" by auto + with A have "M{FPow(X). Q F=0}" by auto + with B have "Q M=0" by auto } moreover { assume "XM" - with A have "M=X" by auto + with A have "M=X" by auto } ultimately have "Q M=0 M=X" by auto } @@ -519,20 +517,20 @@

Theory Topology_ZF_examples

proof- { fix x - assume A:"D X" "X-D ExcludedSet(X,T)" "D0" "xT" "xX" - from A(1) have B:"X-(X-D)=D" by auto - from A(2) have "T(X-D)=0 X-D=X" using ExcludedSet_def by auto + assume A:"D X" "X-D ExcludedSet(X,T)" "D0" "xT" "xX" + from A(1) have B:"X-(X-D)=D" by auto + from A(2) have "T(X-D)=0 X-D=X" using ExcludedSet_def by auto hence "T(X-D)=0 X-(X-D)=X-X" by auto - with B have "T(X-D)=0 D=X-X" by auto + with B have "T(X-D)=0 D=X-X" by auto hence "T(X-D)=0 D=0" by auto - with A(3) have "T(X-D)=0" by auto - with A(4) have "xX-D" by auto - with A(5) have "xD" by auto + with A(3) have "T(X-D)=0" by auto + with A(4) have "xX-D" by auto + with A(5) have "xD" by auto } moreover { - assume A:"XTD" "DX" - from A(1) have "X-DX-(XT)" by auto + assume A:"XTD" "DX" + from A(1) have "X-DX-(XT)" by auto also have " = X-T" by auto finally have "T(X-D) = 0" by auto moreover @@ -545,13 +543,13 @@

Theory Topology_ZF_examples

text‹The interior of a set is itself if it is X› or the difference with the set T› -lemma interior_set_excludedset: +lemma interior_set_excludedset: assumes "AX" shows "Interior(A,ExcludedSet(X,T)) = (if A=X then X else A-T)" proof- { - assume A:"AX" - from assms have "A-T ExcludedSet(X,T)" using ExcludedSet_def by auto + assume A:"AX" + from assms have "A-T ExcludedSet(X,T)" using ExcludedSet_def by auto then have "A-TInterior(A,ExcludedSet(X,T))" using Interior_def by auto moreover @@ -559,7 +557,7 @@

Theory Topology_ZF_examples

fix U assume "U ExcludedSet(X,T)" "UA" then have "TU=0 U=X""UA" using ExcludedSet_def by auto - with A assms have "TU=0""UA" by auto + with A assms have "TU=0""UA" by auto then have "U-T=U""U-TA-T" by auto then have "UA-T" by auto } @@ -577,7 +575,7 @@

Theory Topology_ZF_examples

text‹The closure of a set is itself if it is 0› or the union with T›.› -lemma closure_set_excludedset: +lemma closure_set_excludedset: assumes "AX" shows "Closure(A,ExcludedSet(X,T))=(if A=0 then 0 else A (X T))" proof- @@ -587,14 +585,14 @@

Theory Topology_ZF_examples

hence "Closure(0,ExcludedSet(X,T))=0" by blast moreover { - assume A:"A0" - with assms have "(A(XT)) {is closed in}ExcludedSet(X,T)" using closed_sets_excludedset + assume A:"A0" + with assms have "(A(XT)) {is closed in}ExcludedSet(X,T)" using closed_sets_excludedset by blast then have "(A (X T)) {D Pow(X). D {is closed in}ExcludedSet(X,T) AD}" - using assms by auto + using assms by auto then have "(A (X T))ClosedCovers(A,ExcludedSet(X,T))" unfolding ClosedCovers_def using union_excludedset by auto - then have l1:"ClosedCovers(A,ExcludedSet(X,T)) (A (X T))" by blast + then have l1:"ClosedCovers(A,ExcludedSet(X,T)) (A (X T))" by blast { fix U assume "UClosedCovers(A,ExcludedSet(X,T))" @@ -602,13 +600,13 @@

Theory Topology_ZF_examples

union_excludedset by auto then have "U=0(XT)U" and "AU" using closed_sets_excludedset by auto - with A have "(XT)U""AU" by auto + with A have "(XT)U""AU" by auto hence "(XT)AU" by auto } - with assms have "(A (X T)) ClosedCovers(A,ExcludedSet(X,T))" + with assms have "(A (X T)) ClosedCovers(A,ExcludedSet(X,T))" using topology0.Top_3_L3 topology0_excludedset union_excludedset by auto - with l1 have "ClosedCovers(A,ExcludedSet(X,T)) = (A(XT))" by auto + with l1 have "ClosedCovers(A,ExcludedSet(X,T)) = (A(XT))" by auto then have "Closure(A, ExcludedSet(X,T)) = A(XT)" using Closure_def by auto } @@ -617,7 +615,7 @@

Theory Topology_ZF_examples

text‹The boundary of a set is 0› if $A$ is X› or 0›, and X∩T› in other case. › -lemma boundary_excludedset: +lemma boundary_excludedset: assumes "AX" shows "Boundary(A,ExcludedSet(X,T)) = (if A=0A=X then 0 else XT)" proof- @@ -625,7 +623,7 @@

Theory Topology_ZF_examples

have "Closure(0,ExcludedSet(X,T))=0""Closure(X - 0,ExcludedSet(X,T))=X" using closure_set_excludedset by auto then have "Boundary(0,ExcludedSet(X,T)) = 0"using Boundary_def using - union_excludedset assms by auto + union_excludedset assms by auto } moreover { @@ -638,8 +636,8 @@

Theory Topology_ZF_examples

moreover { assume "A0" and "AX" - then have "X-A0" using assms by auto - with assms A0 AX have "Closure(A,ExcludedSet(X,T)) = A (XT)" + then have "X-A0" using assms by auto + with assms A0 AX have "Closure(A,ExcludedSet(X,T)) = A (XT)" using closure_set_excludedset by simp moreover from AX have "X-A X" by blast @@ -667,13 +665,13 @@

Theory Topology_ZF_examples

text‹If the set which is excluded is disjoint with X›, then the topology is discrete.› -lemma empty_excludedset: +lemma empty_excludedset: assumes "TX=0" shows "ExcludedSet(X,T) = Pow(X)" proof - from assms show "ExcludedSet(X,T) Pow(X)" using smaller_excludedset ExcludedSet_def + from assms show "ExcludedSet(X,T) Pow(X)" using smaller_excludedset ExcludedSet_def by auto - from assms show "Pow(X) ExcludedSet(X,T)" unfolding ExcludedSet_def by blast + from assms show "Pow(X) ExcludedSet(X,T)" unfolding ExcludedSet_def by blast qed text‹The topological subspaces of the ExcludedSet X T› topology @@ -685,10 +683,10 @@

Theory Topology_ZF_examples

{ fix M assume "M(ExcludedSet(X,T) {restricted to} Y)" - then obtain A where A1:"A:ExcludedSet(X,T)" "M=Y A" unfolding RestrictedTo_def by auto + then obtain A where A1:"A:ExcludedSet(X,T)" "M=Y A" unfolding RestrictedTo_def by auto then have "MPow(X Y)" unfolding ExcludedSet_def by auto moreover - from A1 have "TM=0M=YX" unfolding ExcludedSet_def by blast + from A1 have "TM=0M=YX" unfolding ExcludedSet_def by blast ultimately have "M ExcludedSet(Y X,T)" unfolding ExcludedSet_def by auto } @@ -696,7 +694,7 @@

Theory Topology_ZF_examples

{ fix M let ?A = "M ((XY-T)-Y)" - assume A:"M ExcludedSet(YX,T)" + assume A:"M ExcludedSet(YX,T)" { assume "M = Y X" then have "M ExcludedSet(X,T) {restricted to} Y" unfolding RestrictedTo_def @@ -704,19 +702,19 @@

Theory Topology_ZF_examples

} moreover { - assume AS:"MY X" - from A AS have A1:"(MPow(Y X) TM=0)" unfolding ExcludedSet_def by auto + assume AS:"MY X" + from A AS have A1:"(MPow(Y X) TM=0)" unfolding ExcludedSet_def by auto then have "?APow(X)" by blast moreover have "T?A=TM" by blast - with A1 have "T?A=0" by auto + with A1 have "T?A=0" by auto ultimately have "?A ExcludedSet(X,T)" unfolding ExcludedSet_def by auto - then have AT:"Y ?A ExcludedSet(X,T) {restricted to} Y" unfolding RestrictedTo_def + then have AT:"Y ?A ExcludedSet(X,T) {restricted to} Y" unfolding RestrictedTo_def by auto have "Y ?A=Y M" by blast - also have "=M" using A1 by auto + also have "=M" using A1 by auto finally have "Y?A = M" by simp - with AT have "M ExcludedSet(X,T) {restricted to} Y" by auto + with AT have "M ExcludedSet(X,T) {restricted to} Y" by auto } ultimately have "M ExcludedSet(X,T) {restricted to} Y" by auto } @@ -743,10 +741,10 @@

Theory Topology_ZF_examples

{ fix M assume "M Pow(IncludedSet(X,Q))" - then have A:"M{FPow(X). Q F} {0}" using IncludedSet_def by auto + then have A:"M{FPow(X). Q F} {0}" using IncludedSet_def by auto then have "MPow(X)" by auto moreover - have"Q M M=0" using A by blast + have"Q M M=0" using A by blast ultimately have "MIncludedSet(X,Q)" using IncludedSet_def by auto } moreover @@ -783,18 +781,18 @@

Theory Topology_ZF_examples

text‹The topology is defined in the set $X$.› -lemma union_includedset: +lemma union_includedset: assumes "TX " shows "IncludedSet(X,T) = X" proof- - from assms have "X IncludedSet(X,T)" using IncludedSet_def by auto + from assms have "X IncludedSet(X,T)" using IncludedSet_def by auto then show "IncludedSet(X,T) = X" using IncludedSet_def by auto qed text‹The closed sets are those which are disjoint with T› and X›.› -lemma closed_sets_includedset: +lemma closed_sets_includedset: assumes "TX" shows "D {is closed in} IncludedSet(X,T) (DPow(X) (D T)=0) D=X" proof- @@ -802,40 +800,40 @@

Theory Topology_ZF_examples

then have "X-XIncludedSet(X,T)" using IncludedSet_def by auto moreover { - assume A:"D X" "X - D IncludedSet(X,T) "" D X" - from A(2) have "T(X-D) X-D=0" using IncludedSet_def by auto - with A(1) have "T(X-D) D=X" by blast - with A(3) have "T(X-D)" by auto + assume A:"D X" "X - D IncludedSet(X,T) "" D X" + from A(2) have "T(X-D) X-D=0" using IncludedSet_def by auto + with A(1) have "T(X-D) D=X" by blast + with A(3) have "T(X-D)" by auto hence "DT=0" by blast } moreover { - assume A:"DT=0""DX" - from A(1) assms have "T(X-D)" by blast + assume A:"DT=0""DX" + from A(1) assms have "T(X-D)" by blast then have "X-DIncludedSet(X,T)" using IncludedSet_def by auto } - ultimately show ?thesis using IsClosed_def union_includedset assms by auto + ultimately show ?thesis using IsClosed_def union_includedset assms by auto qed text‹The interior of a set is itself if it is open or the empty set if it isn't.› -lemma interior_set_includedset: +lemma interior_set_includedset: assumes "AX" shows "Interior(A,IncludedSet(X,T))= (if TA then A else 0)" proof- { fix x - assume A:"Interior(A,IncludedSet(X,T)) 0 ""xT" + assume A:"Interior(A,IncludedSet(X,T)) 0 ""xT" have "Interior(A,IncludedSet(X,T)) IncludedSet(X,T)" using topology0.Top_2_L2 topology0_includedset by auto - with A(1) have "T Interior(A,IncludedSet(X,T))" using IncludedSet_def + with A(1) have "T Interior(A,IncludedSet(X,T))" using IncludedSet_def by auto - with A(2) have "x Interior(A,IncludedSet(X,T))" by auto + with A(2) have "x Interior(A,IncludedSet(X,T))" by auto then have "xA" using topology0.Top_2_L1 topology0_includedset by auto} moreover { assume "TA" - with assms have "AIncludedSet(X,T)" using IncludedSet_def by auto + with assms have "AIncludedSet(X,T)" using IncludedSet_def by auto then have "Interior(A,IncludedSet(X,T)) = A" using topology0.Top_2_L3 topology0_includedset by auto } @@ -844,22 +842,22 @@

Theory Topology_ZF_examples

text‹The closure of a set is itself if it is closed or the whole space if it is not.› -lemma closure_set_includedset: +lemma closure_set_includedset: assumes "AX" "TX" shows "Closure(A,IncludedSet(X,T)) = (if TA=0 then A else X)" proof- { assume AS:"TA=0" then have "A {is closed in} IncludedSet(X,T)" using closed_sets_includedset - assms by auto - with assms(1) have "Closure(A,IncludedSet(X,T))=A" using topology0.Top_3_L8 - topology0_includedset union_includedset assms(2) by auto + assms by auto + with assms(1) have "Closure(A,IncludedSet(X,T))=A" using topology0.Top_3_L8 + topology0_includedset union_includedset assms(2) by auto } moreover { - assume AS:"TA 0" + assume AS:"TA 0" have "XClosedCovers(A,IncludedSet(X,T))" using ClosedCovers_def - closed_sets_includedset union_includedset assms by auto + closed_sets_includedset union_includedset assms by auto then have l1:"ClosedCovers(A,IncludedSet(X,T))X" using Closure_def by auto moreover @@ -868,13 +866,13 @@

Theory Topology_ZF_examples

assume "UClosedCovers(A,IncludedSet(X,T))" then have "U{is closed in}IncludedSet(X,T)""AU" using ClosedCovers_def by auto - then have "U=X(TU)=0""AU" using closed_sets_includedset assms(2) + then have "U=X(TU)=0""AU" using closed_sets_includedset assms(2) by auto then have "U=X(TA)=0" by auto - then have "U=X" using AS by auto + then have "U=X" using AS by auto } then have "X ClosedCovers(A,IncludedSet(X,T))" using topology0.Top_3_L3 - topology0_includedset union_includedset assms by auto + topology0_includedset union_includedset assms by auto ultimately have "ClosedCovers(A,IncludedSet(X,T))=X" by auto then have "Closure(A,IncludedSet(X,T)) = X " using Closure_def by auto @@ -886,38 +884,38 @@

Theory Topology_ZF_examples

completely and
X› if T› is divided between the two sets. The case where T=0› is considered as a special case.›
-lemma boundary_includedset: +lemma boundary_includedset: assumes "AX" "TX" "T0" shows "Boundary(A,IncludedSet(X,T))=(if TA then X-A else (if TA=0 then A else X))" proof - from AX have "X-A X" by auto { assume "TA" - with assms(2,3) have "TA0" and "T(X-A)=0" by auto - with assms(1,2) X-A X have + with assms(2,3) have "TA0" and "T(X-A)=0" by auto + with assms(1,2) X-A X have "Closure(A,IncludedSet(X,T)) = X" and "Closure(X-A,IncludedSet(X,T)) = (X-A)" using closure_set_includedset by auto - with assms(2) have "Boundary(A,IncludedSet(X,T)) = X-A" + with assms(2) have "Boundary(A,IncludedSet(X,T)) = X-A" using Boundary_def union_includedset by auto } moreover { assume "~(TA)" and "TA=0" - with assms(2) have "T(X-A)0" by auto - with assms(1,2) TA=0 X-A X have + with assms(2) have "T(X-A)0" by auto + with assms(1,2) TA=0 X-A X have "Closure(A,IncludedSet(X,T)) = A" and "Closure(X-A,IncludedSet(X,T)) = X" using closure_set_includedset by auto - with assms(1,2) have "Boundary(A,IncludedSet(X,T))=A" using Boundary_def union_includedset + with assms(1,2) have "Boundary(A,IncludedSet(X,T))=A" using Boundary_def union_includedset by auto } moreover { assume "~(TA)" and "TA 0" - with assms(1,2) have "T(X-A) 0" by auto - with assms(1,2) TA0 X-A X have + with assms(1,2) have "T(X-A) 0" by auto + with assms(1,2) TA0 X-A X have "Closure(A,IncludedSet(X,T)) = X" and "Closure(X-A,IncludedSet(X,T)) = X" using closure_set_includedset by auto - with assms(2) have "Boundary(A,IncludedSet(X,T)) = X" + with assms(2) have "Boundary(A,IncludedSet(X,T)) = X" using Boundary_def union_includedset by auto } ultimately show ?thesis by auto @@ -941,11 +939,11 @@

Theory Topology_ZF_examples

text‹If the set which is included is not a subset of X›, then the topology is trivial.› -lemma empty_includedset: +lemma empty_includedset: assumes "~(TX)" shows "IncludedSet(X,T) = {0}" proof - from assms show "IncludedSet(X,T) {0}" and "{0} IncludedSet(X,T)" + from assms show "IncludedSet(X,T) {0}" and "{0} IncludedSet(X,T)" unfolding IncludedSet_def by auto qed @@ -955,18 +953,18 @@

Theory Topology_ZF_examples

is never trivial. There is no need for a separate proof because the only subspace of the trivial topology is itself.›
-lemma subspace_includedset: +lemma subspace_includedset: assumes "TX" shows "IncludedSet(X,T) {restricted to} Y = IncludedSet(YX,YT)" proof { fix M assume "M (IncludedSet(X,T) {restricted to} Y)" - then obtain A where A1:"A:IncludedSet(X,T)" "M = YA" unfolding RestrictedTo_def + then obtain A where A1:"A:IncludedSet(X,T)" "M = YA" unfolding RestrictedTo_def by auto then have "M Pow(XY)" unfolding IncludedSet_def by auto moreover - from A1 have "YTM M=0" unfolding IncludedSet_def by blast + from A1 have "YTM M=0" unfolding IncludedSet_def by blast ultimately have "M IncludedSet(YX, YT)" unfolding IncludedSet_def by auto } @@ -975,7 +973,7 @@

Theory Topology_ZF_examples

{ fix M let ?A = "M T" - assume A:"M IncludedSet(YX, YT)" + assume A:"M IncludedSet(YX, YT)" { assume "M=0" then have "MIncludedSet(X,T) {restricted to} Y" unfolding RestrictedTo_def @@ -983,18 +981,18 @@

Theory Topology_ZF_examples

} moreover { - assume AS:"M0" - from A AS have A1:"MPow(YX) YTM" unfolding IncludedSet_def by auto - then have "?APow(X)" using assms by blast + assume AS:"M0" + from A AS have A1:"MPow(YX) YTM" unfolding IncludedSet_def by auto + then have "?APow(X)" using assms by blast moreover have "T?A" by blast ultimately have "?A IncludedSet(X,T)" unfolding IncludedSet_def by auto - then have AT:"Y ?A IncludedSet(X,T) {restricted to} Y"unfolding RestrictedTo_def + then have AT:"Y ?A IncludedSet(X,T) {restricted to} Y"unfolding RestrictedTo_def by auto - from A1 have "Y ?A=Y M" by blast - also from A1 have "=M" by auto + from A1 have "Y ?A=Y M" by blast + also from A1 have "=M" by auto finally have "Y?A = M" by simp - with AT have "M IncludedSet(X,T) {restricted to} Y" + with AT have "M IncludedSet(X,T) {restricted to} Y" by auto } ultimately have "M IncludedSet(X,T) {restricted to} Y" by auto diff --git a/docs/IsarMathLib/UniformSpace_ZF.html b/docs/IsarMathLib/UniformSpace_ZF.html index 93b0402..ada6cfa 100644 --- a/docs/IsarMathLib/UniformSpace_ZF.html +++ b/docs/IsarMathLib/UniformSpace_ZF.html @@ -47,7 +47,7 @@

Theory UniformSpace_ZF

text‹ This theory defines uniform spaces and proves their basic properties. › -subsection‹ Definition and motivation › +subsection‹ Entourages and neighborhoods › text‹ Just like a topological space constitutes the minimal setting in which one can speak of continuous functions, the notion of uniform spaces @@ -70,29 +70,35 @@

Theory UniformSpace_ZF

"Φ {is a uniformity on} X (Φ {is a filter on} (X×X)) (UΦ. id(X) U (VΦ. V O V U) converse(U) Φ)" +text‹Since the whole $X\times X$ is in a uniformity, a uniformity is never empty.› + +lemma uniformity_non_empty: assumes "Φ {is a uniformity on} X" + shows "Φ" + using assms unfolding IsUniformity_def IsFilter_def by auto + text‹ If $\Phi$ is a uniformity on $X$, then the every element $V$ of $\Phi$ is a certain relation on $X$ (a subset of $X\times X$) and is called an ''entourage''. For an $x\in X$ we call $V\{ x\}$ a neighborhood of $x$. The first useful fact we will show is that neighborhoods are non-empty. › -lemma neigh_not_empty: +lemma neigh_not_empty: assumes "Φ {is a uniformity on} X" "WΦ" and "xX" - shows "W``{x} 0" and "x W``{x}" + shows "W``{x} " and "x W``{x}" proof - - from assms(1,2) have "id(X) W" + from assms(1,2) have "id(X) W" unfolding IsUniformity_def IsFilter_def by auto - with xX show" xW``{x}" and "W``{x} 0" by auto + with xX show" xW``{x}" and "W``{x} " by auto qed text‹The filter part of the definition of uniformity for easier reference:› -lemma unif_filter: assumes "Φ {is a uniformity on} X" +lemma unif_filter: assumes "Φ {is a uniformity on} X" shows "Φ {is a filter on} (X×X)" - using assms unfolding IsUniformity_def by simp + using assms unfolding IsUniformity_def by simp text‹The second part of the definition of uniformity for easy reference:› -lemma entourage_props: +lemma entourage_props: assumes "Φ {is a uniformity on} X" and "AΦ" shows "A X×X" @@ -100,9 +106,9 @@

Theory UniformSpace_ZF

"VΦ. V O V A" "converse(A) Φ" proof - - from assms show "id(X) A" "VΦ. V O V A" "converse(A) Φ" + from assms show "id(X) A" "VΦ. V O V A" "converse(A) Φ" unfolding IsUniformity_def by auto - from assms show "A X×X" + from assms show "A X×X" using unif_filter unfolding IsFilter_def by blast qed @@ -110,13 +116,13 @@

Theory UniformSpace_ZF

of uniformity $\Phi$ there is another one, say $V$ such that $V\circ V\subseteq U$. Sometimes such $V$ is said to be half the size of $U$. The next lemma states that $V$ can be taken to be symmetric. ›
-lemma half_size_symm: assumes "Φ {is a uniformity on} X" "WΦ" +lemma half_size_symm: assumes "Φ {is a uniformity on} X" "WΦ" shows "VΦ. V O V W V=converse(V)" proof - - from assms obtain U where "UΦ" and "U O U W" + from assms obtain U where "UΦ" and "U O U W" unfolding IsUniformity_def by auto let ?V = "U converse(U)" - from assms(1) UΦ have "?V Φ" and "?V = converse(?V)" + from assms(1) UΦ have "?V Φ" and "?V = converse(?V)" unfolding IsUniformity_def IsFilter_def by auto moreover from U O U W have "?V O ?V W" by auto ultimately show ?thesis by blast @@ -125,16 +131,16 @@

Theory UniformSpace_ZF

text‹Inside every member $W$ of the uniformity $\Phi$ we can find one that is symmetric and smaller than a third of size $W$. Compare with the Metamath's theorem with the same name.› -lemma ustex3sym: assumes "Φ {is a uniformity on} X" "AΦ" +lemma ustex3sym: assumes "Φ {is a uniformity on} X" "AΦ" shows "BΦ. B O (B O B) A B=converse(B)" proof - - from assms obtain C where "CΦ" and "C O C A" + from assms obtain C where "CΦ" and "C O C A" unfolding IsUniformity_def by auto - from assms(1) CΦ obtain B where + from assms(1) CΦ obtain B where "BΦ" "B O B C" and "B=converse(B)" using half_size_symm by blast with C O C A have "(B O B) O (B O B) A" by blast - with assms(1) BΦ have "B O (B O B) A" + with assms(1) BΦ have "B O (B O B) A" using entourage_props(1,2) by blast with BΦ B=converse(B) show ?thesis by blast qed @@ -142,17 +148,17 @@

Theory UniformSpace_ZF

text‹If $\Phi$ is a uniformity on $X$ then every element of $\Phi$ is a subset of $X\times X$ whose domain is $X$. › -lemma uni_domain: +lemma uni_domain: assumes "Φ {is a uniformity on} X" "WΦ" shows "W X×X" and "domain(W) = X" proof - - from assms show "W X×X" unfolding IsUniformity_def IsFilter_def + from assms show "W X×X" unfolding IsUniformity_def IsFilter_def by blast show "domain(W) = X" proof - from assms show "domain(W) X" unfolding IsUniformity_def IsFilter_def + from assms show "domain(W) X" unfolding IsUniformity_def IsFilter_def by auto - from assms show "X domain(W)" unfolding IsUniformity_def by blast + from assms show "X domain(W)" unfolding IsUniformity_def by blast qed qed @@ -160,10 +166,10 @@

Theory UniformSpace_ZF

the image of the singleton $\{ x\}$ by $W$ is contained in $X$. Compare the Metamath's theorem with the same name. ›
-lemma ustimasn: +lemma ustimasn: assumes "Φ {is a uniformity on} X" "WΦ" and "xX" shows "W``{x} X" - using assms uni_domain(1) by auto + using assms uni_domain(1) by auto text‹ Uniformity Φ› defines a natural topology on its space $X$ via the neighborhood system that assigns the collection $\{V(\{x\}):V\in \Phi\}$ to every point $x\in X$. @@ -172,38 +178,38 @@

Theory UniformSpace_ZF

fact which is useful to shorten the remaining proofs, usually treated as obvious in standard mathematics. ›
-lemma neigh_filt_fun: +lemma neigh_filt_fun: assumes "Φ {is a uniformity on} X" defines " {x,{V``{x}.VΦ}.xX}" shows ":XPow(Pow(X))" and "xX. `(x) = {V``{x}.VΦ}" proof - - from assms have "xX. {V``{x}.VΦ} Pow(Pow(X))" + from assms have "xX. {V``{x}.VΦ} Pow(Pow(X))" using IsUniformity_def IsFilter_def image_subset by auto - with assms show ":XPow(Pow(X))" using ZF_fun_from_total by simp - with assms show "xX. `(x) = {V``{x}.VΦ}" using ZF_fun_from_tot_val + with assms show ":XPow(Pow(X))" using ZF_fun_from_total by simp + with assms show "xX. `(x) = {V``{x}.VΦ}" using ZF_fun_from_tot_val by simp qed text‹ In the next lemma we show that the collection defined in lemma neigh_filt_fun› is a filter on $X$. The proof is kind of long, but it just checks that all filter conditions hold.› -lemma filter_from_uniformity: +lemma filter_from_uniformity: assumes "Φ {is a uniformity on} X" and "xX" defines " {x,{V``{x}.VΦ}.xX}" shows "`(x) {is a filter on} X" proof - - from assms have PhiFilter: "Φ {is a filter on} (X×X)" and + from assms have PhiFilter: "Φ {is a filter on} (X×X)" and ":XPow(Pow(X))" and "`(x) = {V``{x}.VΦ}" using IsUniformity_def neigh_filt_fun by auto - have "0 `(x)" + have " `(x)" proof - - from assms xX have "0 {V``{x}.VΦ}" using neigh_not_empty by blast - with `(x) = {V``{x}.VΦ} show "0 `(x)" by simp + from assms xX have " {V``{x}.VΦ}" using neigh_not_empty by blast + with `(x) = {V``{x}.VΦ} show " `(x)" by simp qed moreover have "X `(x)" proof - note `(x) = {V``{x}.VΦ} - moreover from assms have "X×X Φ" unfolding IsUniformity_def IsFilter_def + moreover from assms have "X×X Φ" unfolding IsUniformity_def IsFilter_def by blast hence "(X×X)``{x} {V``{x}.VΦ}" by auto moreover from xX have "(X×X)``{x} = X" by auto @@ -211,17 +217,17 @@

Theory UniformSpace_ZF

qed moreover from :XPow(Pow(X)) xX have "`(x) Pow(X)" using apply_funtype by blast - moreover have LargerIn: "B `(x). C Pow(X). BC C `(x)" + moreover have LargerIn: "B `(x). C Pow(X). BC C `(x)" proof - { fix B assume "B `(x)" fix C assume "C Pow(X)" and "BC" from `(x) = {V``{x}.VΦ} B `(x) obtain U where "UΦ" and "B = U``{x}" by auto let ?V = "U C×C" - from assms UΦ C Pow(X) have "?V Pow(X×X)" and "U?V" + from assms UΦ C Pow(X) have "?V Pow(X×X)" and "U?V" using IsUniformity_def IsFilter_def by auto - with UΦ PhiFilter have "?VΦ" using IsFilter_def by simp - moreover from assms UΦ xX B = U``{x} BC have "C = ?V``{x}" + with UΦ PhiFilter have "?VΦ" using IsFilter_def by simp + moreover from assms UΦ xX B = U``{x} BC have "C = ?V``{x}" using neigh_not_empty image_greater_rel by simp ultimately have "C {V``{x}.VΦ}" by auto with `(x) = {V``{x}.VΦ} have "C `(x)" by simp @@ -234,13 +240,13 @@

Theory UniformSpace_ZF

"A = VA``{x}" "B = VB``{x}" and "VA Φ" "VB Φ" by auto let ?C = "VA``{x} VB``{x}" - from assms VA Φ VB Φ have "VAVB Φ" using IsUniformity_def IsFilter_def + from assms VA Φ VB Φ have "VAVB Φ" using IsUniformity_def IsFilter_def by simp with `(x) = {V``{x}.VΦ} have "(VAVB)``{x} `(x)" by auto - moreover from PhiFilter VA Φ VB Φ have "?C Pow(X)" unfolding IsFilter_def + moreover from PhiFilter VA Φ VB Φ have "?C Pow(X)" unfolding IsFilter_def by auto moreover have "(VAVB)``{x} ?C" using image_Int_subset_left by simp - moreover note LargerIn + moreover note LargerIn ultimately have "?C `(x)" by simp with A = VA``{x} B = VB``{x} have "AB `(x)" by blast } thus ?thesis by simp @@ -251,20 +257,20 @@

Theory UniformSpace_ZF

text‹A rephrasing of filter_from_uniformity›: if $\Phi$ is a uniformity on $X$, then $\{V(\{ x\}) | V\in \Phi\}$ is a filter on $X$ for every $x\in X$.› -lemma unif_filter_at_point: +lemma unif_filter_at_point: assumes "Φ {is a uniformity on} X" and "xX" shows "{V``{x}.VΦ} {is a filter on} X" - using assms filter_from_uniformity ZF_fun_from_tot_val1 + using assms filter_from_uniformity ZF_fun_from_tot_val1 by simp text‹A frequently used property of filters is that they are "upward closed" i.e. supersets of a filter element are also in the filter. The next lemma makes this explicit for easy reference as applied to the natural filter created from a uniformity.› -corollary unif_filter_up_closed: +corollary unif_filter_up_closed: assumes "Φ {is a uniformity on} X" "xX" "U {V``{x}. VΦ}" "WX" "UW" shows "W {V``{x}.VΦ}" - using assms filter_from_uniformity ZF_fun_from_tot_val1 + using assms filter_from_uniformity ZF_fun_from_tot_val1 unfolding IsFilter_def by auto text‹ The function defined in the premises of lemma neigh_filt_fun› @@ -272,14 +278,14 @@

Theory UniformSpace_ZF

of the "half-the-size" neighborhood condition (
∃V∈Φ. V O V ⊆ U›) of the uniformity definition, but not the converse(U) ∈ Φ› part. ›
-theorem neigh_from_uniformity: +theorem neigh_from_uniformity: assumes "Φ {is a uniformity on} X" shows "{x,{V``{x}.VΦ}.xX} {is a neighborhood system on} X" proof - let ?ℳ = "{x,{V``{x}.VΦ}.xX}" - from assms have "?ℳ:XPow(Pow(X))" and Mval: "xX. ?ℳ`(x) = {V``{x}.VΦ}" + from assms have "?ℳ:XPow(Pow(X))" and Mval: "xX. ?ℳ`(x) = {V``{x}.VΦ}" using IsUniformity_def neigh_filt_fun by auto - moreover from assms have "xX. (?ℳ`(x) {is a filter on} X)" using filter_from_uniformity + moreover from assms have "xX. (?ℳ`(x) {is a filter on} X)" using filter_from_uniformity by simp moreover { fix x assume "xX" @@ -288,21 +294,21 @@

Theory UniformSpace_ZF

{ fix N assume "N?ℳ`(x)" have "xN" and "U?ℳ`(x).yU.(N ?ℳ`(y))" proof - - from ?ℳ:XPow(Pow(X)) Mval xX N?ℳ`(x) + from ?ℳ:XPow(Pow(X)) Mval xX N?ℳ`(x) obtain U where "UΦ" and "N = U``{x}" by auto - with assms xX show "xN" using neigh_not_empty by simp - from assms UΦ obtain V where "VΦ" and "V O V U" + with assms xX show "xN" using neigh_not_empty by simp + from assms UΦ obtain V where "VΦ" and "V O V U" unfolding IsUniformity_def by auto let ?W = "V``{x}" - from VΦ Mval xX have "?W ?ℳ`(x)" by auto + from VΦ Mval xX have "?W ?ℳ`(x)" by auto moreover have "y?W. N ?ℳ`(y)" proof - { fix y assume "y?W" with ?ℳ:XPow(Pow(X)) xX ?W ?ℳ`(x) have "yX" using apply_funtype by blast - with assms have "?ℳ`(y) {is a filter on} X" using filter_from_uniformity + with assms have "?ℳ`(y) {is a filter on} X" using filter_from_uniformity by simp - moreover from assms yX VΦ have "V``{y} ?ℳ`(y)" + moreover from assms yX VΦ have "V``{y} ?ℳ`(y)" using neigh_filt_fun by auto moreover from ?ℳ:XPow(Pow(X)) xX N ?ℳ`(x) have "N Pow(X)" using apply_funtype by blast @@ -347,11 +353,11 @@

Theory UniformSpace_ZF

text‹ The collection of sets constructed in the UniformTopology › definition is indeed a topology on $X$. › -theorem uniform_top_is_top: +theorem uniform_top_is_top: assumes "Φ {is a uniformity on} X" shows "UniformTopology(Φ,X) {is a topology}" and " UniformTopology(Φ,X) = X" - using assms neigh_from_uniformity uniftop_def_alt topology_from_neighs + using assms neigh_from_uniformity uniftop_def_alt topology_from_neighs by auto text‹If we have a uniformity $\Phi$ we can create a neighborhood system from it in two ways. @@ -362,27 +368,27 @@

Theory UniformSpace_ZF

as in theorem
neigh_from_topology›. The next theorem states that these two ways give the same result. ›
-theorem neigh_unif_same: assumes "Φ {is a uniformity on} X" +theorem neigh_unif_same: assumes "Φ {is a uniformity on} X" shows "{x,{V``{x}.VΦ}. xX} = {neighborhood system of} UniformTopology(Φ,X)" - using assms neigh_from_uniformity nei_top_nei_round_trip uniftop_def_alt + using assms neigh_from_uniformity nei_top_nei_round_trip uniftop_def_alt by simp text‹Another form of the definition of topology generated from a uniformity.› -lemma uniftop_def_alt1: assumes "Φ {is a uniformity on} X" +lemma uniftop_def_alt1: assumes "Φ {is a uniformity on} X" shows "UniformTopology(Φ,X) = {UPow(X). xU. WΦ. W``{x} U}" proof let ?T = "UniformTopology(Φ,X)" show "?T {UPow(X). xU. WΦ. W``{x} U}" unfolding UniformTopology_def by auto { fix U assume "U{UPow(X). xU. WΦ. W``{x} U}" - then have "UPow(X)" and I: "xU. WΦ. W``{x} U" by auto + then have "UPow(X)" and I: "xU. WΦ. W``{x} U" by auto { fix x assume "xU" - with I obtain W where "WΦ" and "W``{x} U" + with I obtain W where "WΦ" and "W``{x} U" by auto let ?𝔉 = "{V``{x}.VΦ}" - from assms(1) UPow(X) xU WΦ have + from assms(1) UPow(X) xU WΦ have "?𝔉 {is a filter on} X" and "W``{x} ?𝔉" using unif_filter_at_point by auto with UPow(X) W``{x} U have "U?𝔉" @@ -396,14 +402,14 @@

Theory UniformSpace_ZF

text‹Images of singletons by entourages are neighborhoods of those singletons.› -lemma image_singleton_ent_nei: +lemma image_singleton_ent_nei: assumes "Φ {is a uniformity on} X" "VΦ" "xX" defines " {neighborhood system of} UniformTopology(Φ,X)" shows "V``{x} `(x)" proof - - from assms(1,4) have " = {x,{V``{x}.VΦ}. xX}" + from assms(1,4) have " = {x,{V``{x}.VΦ}. xX}" using neigh_unif_same by simp - with assms(2,3) show ?thesis + with assms(2,3) show ?thesis using ZF_fun_from_tot_val1 by auto qed @@ -411,20 +417,20 @@

Theory UniformSpace_ZF

of images of the singleton by the entourages $W\in \Phi$. See also the Metamath's theorem with the same name. ›
-lemma utopsnneip: assumes "Φ {is a uniformity on} X" "xX" +lemma utopsnneip: assumes "Φ {is a uniformity on} X" "xX" defines "𝒮 {set neighborhood system of} UniformTopology(Φ,X)" shows "𝒮`{x} = {W``{x}. WΦ}" proof - let ?T = "UniformTopology(Φ,X)" let ?ℳ = "{neighborhood system of} ?T" - from assms(1,2) have "x ?T" + from assms(1,2) have "x ?T" using uniform_top_is_top(2) by simp - with assms(3) have "?ℳ`(x) = 𝒮`{x}" + with assms(3) have "?ℳ`(x) = 𝒮`{x}" using neigh_from_nei by simp moreover - from assms(1) have "?ℳ = {x,{W``{x}.WΦ}. xX}" + from assms(1) have "?ℳ = {x,{W``{x}.WΦ}. xX}" using neigh_unif_same by simp - with assms(2) have "?ℳ`(x) = {W``{x}.WΦ}" + with assms(2) have "?ℳ`(x) = {W``{x}.WΦ}" using ZF_fun_from_tot_val1 by simp ultimately show ?thesis by simp qed @@ -432,9 +438,9 @@

Theory UniformSpace_ZF

text‹Images of singletons by entourages are set neighborhoods of those singletons. See also the Metamath theorem with the same name.› -corollary utopsnnei: assumes "Φ {is a uniformity on} X" "WΦ" "xX" +corollary utopsnnei: assumes "Φ {is a uniformity on} X" "WΦ" "xX" defines "𝒮 {set neighborhood system of} UniformTopology(Φ,X)" - shows "W``{x} 𝒮`{x}" using assms utopsnneip by auto + shows "W``{x} 𝒮`{x}" using assms utopsnneip by auto text‹If $\Phi$ is a uniformity on $X$ that generates a topology $T$, $R$ is any relation on $X$ (i.e. $R\subseteq X\times X$), $W$ is a symmetric entourage (i.e. $W\in \Phi$, @@ -442,33 +448,33 @@

Theory UniformSpace_ZF

is contained the the composition $V\circ (M \circ V)$. Metamath has a similar theorem with the same name. ›
-lemma utop3cls: +lemma utop3cls: assumes "Φ {is a uniformity on} X" "RX×X" "WΦ" "W=converse(W)" defines "J UniformTopology(Φ,X)" shows "Closure(R,J×tJ) W O (R O W)" proof let ?M = "{set neighborhood system of} (J×tJ)" - fix z assume zMem: "z Closure(R,J×tJ)" - from assms(1,5) have Jtop: "J {is a topology}" and "J = X" + fix z assume zMem: "z Closure(R,J×tJ)" + from assms(1,5) have Jtop: "J {is a topology}" and "J = X" using uniform_top_is_top by auto - then have JJtop: "(J×tJ) {is a topology}" and JxJ: "(J×tJ) = X×X" + then have JJtop: "(J×tJ) {is a topology}" and JxJ: "(J×tJ) = X×X" using Top_1_4_T1(1,3) by auto - with assms(2) have "topology0(J×tJ)" and "R (J×tJ)" + with assms(2) have "topology0(J×tJ)" and "R (J×tJ)" unfolding topology0_def by auto then have "Closure(R,J×tJ) (J×tJ)" using topology0.Top_3_L11(1) by simp - with z Closure(R,J×tJ) JxJ have "zX×X" by auto + with z Closure(R,J×tJ) JxJ have "zX×X" by auto let ?x = "fst(z)" let ?y = "snd(z)" from zX×X have "?xX" "?yX" "z = ?x,?y" by auto - with assms(1,3,5) Jtop have "(W``{?x})×(W``{?y}) ?M`({?x}×{?y})" + with assms(1,3,5) Jtop have "(W``{?x})×(W``{?y}) ?M`({?x}×{?y})" using utopsnnei neitx by simp moreover from z = ?x,?y have "{?x}×{?y} = {z}" by (rule pair_prod) ultimately have "(W``{?x})×(W``{?y}) ?M`{z}" by simp - with zMem JJtop R (J×tJ) have "(W``{?x})×(W``{?y}) R 0" + with zMem JJtop R (J×tJ) have "(W``{?x})×(W``{?y}) R " using neindisj by blast - with assms(4) have "?x,?y W O (R O W)" + with assms(4) have "?x,?y W O (R O W)" using sym_rel_comp by simp with z = ?x,?y show "z W O (R O W)" by simp @@ -476,14 +482,14 @@

Theory UniformSpace_ZF

text‹Uniform spaces are regular ($T_3$). › -theorem utopreg: +theorem utopreg: assumes "Φ {is a uniformity on} X" shows "UniformTopology(Φ,X) {is regular}" proof - let ?J = "UniformTopology(Φ,X)" let ?𝒮 = "{set neighborhood system of} ?J" - from assms have "?J = X" - and Jtop: "?J {is a topology}" and cntx: "topology0(?J)" + from assms have "?J = X" + and Jtop: "?J {is a topology}" and cntx: "topology0(?J)" using uniform_top_is_top unfolding topology0_def by auto have "U?J. xU. V?J. xV Closure(V,?J)U" proof - @@ -491,13 +497,13 @@

Theory UniformSpace_ZF

then have "U ?𝒮`{x}" using open_nei_singl by simp from U?J have "U?J" by auto with xU ?J = X have "xX" by auto - from assms(1) xX U ?𝒮`{x} obtain A + from assms(1) xX U ?𝒮`{x} obtain A where "U=A``{x}" and "AΦ" using utopsnneip by auto - from assms(1) AΦ obtain W where - "WΦ" "W O (W O W) A" and Wsymm: "W=converse(W)" + from assms(1) AΦ obtain W where + "WΦ" "W O (W O W) A" and Wsymm: "W=converse(W)" using ustex3sym by blast - with assms(1) xX have "W``{x} ?𝒮`{x}" and "W``{x} X" + with assms(1) xX have "W``{x} ?𝒮`{x}" and "W``{x} X" using utopsnnei ustimasn by auto from W``{x} ?𝒮`{x} have "V?J. {x}V VW``{x}" by (rule neii2) @@ -505,9 +511,9 @@

Theory UniformSpace_ZF

by blast have "Closure(V,?J) U" proof - - from assms(1) WΦ ?J = X have "W X×X" + from assms(1) WΦ ?J = X have "W X×X" using entourage_props(1) by simp - from cntx W``{x} X ?J = X VW``{x} + from cntx W``{x} X ?J = X VW``{x} have "Closure(V,?J) Closure(W``{x},?J)" using topology0.top_closure_mono by simp also have "Closure(W``{x},?J) Closure(W,?J×t?J)``{x}" @@ -517,7 +523,7 @@

Theory UniformSpace_ZF

with ?J {is a topology} show ?thesis using imasncls by simp qed - also from assms(1) WX×X WΦ Wsymm W O (W O W) A + also from assms(1) WX×X WΦ Wsymm W O (W O W) A have "Closure(W,?J×t?J)``{x} A``{x}" using utop3cls by blast finally have "Closure(V,?J) A``{x}" @@ -528,10 +534,229 @@

Theory UniformSpace_ZF

by blast } thus ?thesis by simp qed - with Jtop show "?J {is regular}" using is_regular_def_alt + with Jtop show "?J {is regular}" using is_regular_def_alt + by simp +qed + +subsection‹ Base of a uniformity › + +text‹A base› or a fundamental system of entourages› of a uniformity $\Phi$ is + a subset of $\Phi$ that is sufficient to uniquely determine it. This is + analogous to the notion of a base of a topology (see Topology_ZF_1› or a base of a filter + (see Topology_ZF_4›). › + +text‹A base of a uniformity $\Phi$ is any subset $\mathfrak{B}\subseteq \Phi$ such that + every entourage in $\Phi$ contains (at least) one from $\mathfrak{B}$. + The phrase is a base for› is already defined to mean a base for a topology, + so we use the phrase is a uniform base of› here. › + +definition + IsUniformityBase ("_ {is a uniform base of} _" 90) where + "𝔅 {is a uniform base of} Φ 𝔅 Φ (UΦ. B𝔅. BU)" + +text‹Symmetric entourages form a base of the uniformity.› + +lemma symm_are_base: assumes "Φ {is a uniformity on} X" + shows "{VΦ. V = converse(V)} {is a uniform base of} Φ" +proof - + let ?𝔅 = "{VΦ. V = converse(V)}" + { fix W assume "WΦ" + with assms obtain V where "VΦ" "V O V W" "V=converse(V)" + using half_size_symm by blast + from assms VΦ have "V V O V" + using entourage_props(1,2) refl_square_greater by blast + with V O V W VΦ V=converse(V) have "V?𝔅. VW" by auto + } hence "WΦ. V?𝔅. V W" by auto + then show ?thesis unfolding IsUniformityBase_def by auto +qed + +text‹Given a base of a uniformity we can recover the uniformity taking the supersets. + The Supersets› constructor is defined in ZF1›.› + +lemma uniformity_from_base: + assumes "Φ {is a uniformity on} X" "𝔅 {is a uniform base of} Φ" + shows "Φ = Supersets(X×X,𝔅)" +proof + from assms show "Φ Supersets(X×X,𝔅)" + unfolding IsUniformityBase_def Supersets_def + using entourage_props(1) by auto + from assms show "Supersets(X×X,𝔅) Φ" + unfolding Supersets_def IsUniformityBase_def IsUniformity_def IsFilter_def + by auto +qed + +text‹Analogous to the predicate "satisfies base condition" (defined in Topology_ZF_1›) + and "is a base filter" (defined in Topology_ZF_4›) we can specify conditions + for a collection $\mathfrak{B}$ of subsets of $X\times X$ to be a base of some + uniformity on $X$. Namely, the following conditions are necessary and sufficient: + + 1. Intersection of two sets of $\mathfrak{B}$ contains a set of $\mathfrak{B}$. + + 2. Every set of $\mathfrak{B}$ contains the diagonal of $X\times X$. + + 3. For each set $B_1\in \mathfrak{B}$ we can find a set $B_2\in \mathfrak{B}$ + such that $B_2\subseteq B_1^{-1}$. + + 4. For each set $B_1\in \mathfrak{B}$ we can find a set $B_2\in \mathfrak{B}$ + such that $B_2\circ B_2 \subseteq B_1$. + + The conditions are taken from + N. Bourbaki "Elements of Mathematics, General Topology", Chapter II.1., + except for the last two that are missing there.› + +definition + IsUniformityBaseOn ("_ {is a uniform base on} _" 90) where + "𝔅 {is a uniform base on} X + (B1𝔅. B2𝔅. B3𝔅. B3B1B2) (B𝔅. id(X)B) + (B1𝔅. B2𝔅. B2 converse(B1)) (B1𝔅. B2𝔅. B2 O B2 B1) + 𝔅Pow(X×X) 𝔅" + +text‹The next lemma splits the definition of IsUniformityBaseOn› into four conditions + to enable more precise references in proofs.› + +lemma uniformity_base_props: assumes "𝔅 {is a uniform base on} X" + shows + "B1𝔅. B2𝔅. B3𝔅. B3B1B2" + "B𝔅. id(X)B" + "B1𝔅. B2𝔅. B2 converse(B1)" + "B1𝔅. B2𝔅. B2 O B2 B1" + "𝔅Pow(X×X)" and "𝔅" + using assms unfolding IsUniformityBaseOn_def by simp_all + +text‹If supersets of some collection of subsets of $X\times X$ form a uniformity, + then this collection satisfies the conditions in the definition of IsUniformityBaseOn›. › + +theorem base_is_uniform_base: + assumes "𝔅Pow(X×X)" and "Supersets(X×X,𝔅) {is a uniformity on} X" + shows "𝔅 {is a uniform base on} X" +proof - + let = "Supersets(X×X,𝔅)" + have "B1𝔅. B2𝔅. B3𝔅. B3B1B2" + proof - + { fix B1 B2 assume "B1𝔅" "B2𝔅" + with assms(1) have "B1" and "B2" unfolding Supersets_def by auto + with assms(2) have "B3𝔅. B3 B1B2" + unfolding IsUniformity_def IsFilter_def Supersets_def by simp + } thus ?thesis by simp + qed + moreover have "B𝔅. id(X)B" + proof - + { fix B assume "B𝔅" + with assms(1) have "B" unfolding Supersets_def by auto + with assms(2) have "id(X)B" unfolding IsUniformity_def by simp + } thus ?thesis by simp + qed + moreover have "B1𝔅. B2𝔅. B2 converse(B1)" + proof - + { fix B1 assume "B1𝔅" + with assms(1) have "B1" unfolding Supersets_def by auto + with assms have "B2𝔅. B2 converse(B1)" + unfolding IsUniformity_def Supersets_def by auto + } thus ?thesis by simp + qed + moreover have "B1𝔅. B2𝔅. B2 O B2 B1" + proof - + { fix B1 assume "B1𝔅" + with assms(1) have "B1" unfolding Supersets_def by auto + with assms(2) obtain V where "V" and "V O V B1" + unfolding IsUniformity_def by blast + from assms(2) V obtain B2 where "B2𝔅" and "B2V" + unfolding Supersets_def by auto + from V O V B1 B2V have "B2 O B2 B1" by auto + with B2𝔅 have "B2𝔅. B2 O B2 B1" by auto + } thus ?thesis by simp + qed + moreover from assms(2) have "𝔅" + using supersets_of_empty uniformity_non_empty by blast + ultimately show "𝔅 {is a uniform base on} X" + unfolding IsUniformityBaseOn_def using assms(1) by simp +qed + +text‹if a nonempty collection of subsets of $X\times X$ satisfies conditions in the definition + of IsUniformityBaseOn› then the supersets of that collection form a uniformity on $X$.› + +theorem uniformity_base_is_base: + assumes "X" and "𝔅 {is a uniform base on} X" + shows "Supersets(X×X,𝔅) {is a uniformity on} X" +proof - + let = "Supersets(X×X,𝔅)" + from assms(2) have "𝔅Pow(X×X)" using uniformity_base_props(5) by simp + have " {is a filter on} (X×X)" + proof - + from assms have "" + unfolding Supersets_def using uniformity_base_props(2) + by blast + moreover have "X×X " + proof - + from assms(2) obtain B where "B𝔅" + using uniformity_base_props(6) by blast + with 𝔅Pow(X×X) show "X×X " unfolding Supersets_def + by blast + qed + moreover have " Pow(X×X)" unfolding Supersets_def by auto + moreover have "U. V. UV " + proof - + { fix U V assume "U" "V" + then obtain B1 B2 where "B1𝔅" "B2𝔅" "B1U" "B2V" + unfolding Supersets_def by auto + from assms(2) B1𝔅 B2𝔅 obtain B3 where "B3𝔅" and "B3B1B2" + using uniformity_base_props(1) by blast + from B1U B2V B3B1B2 have "B3UV" by auto + with U V B3𝔅 have "UV " + unfolding Supersets_def by auto + } thus ?thesis by simp + qed + moreover have "U. CPow(X×X). UC C" + proof - + { fix U C assume "U" "CPow(X×X)" "UC" + from U obtain B where "B𝔅" and "BU" + unfolding Supersets_def by auto + with UC CPow(X×X) have "C" + unfolding Supersets_def by auto + } thus ?thesis by auto + qed + ultimately show " {is a filter on} (X×X)" + unfolding IsFilter_def by simp + qed + moreover have "U. id(X) U (V. V O V U) converse(U) " + proof - + { fix U assume "U" + then obtain B where "B𝔅" and "BU" + unfolding Supersets_def by auto + with assms(2) have "id(X) U" + using uniformity_base_props(2) by blast + moreover + from assms(2) B𝔅 obtain V where "V𝔅" and "V O V B" + using uniformity_base_props(4) by blast + with 𝔅Pow(X×X) have "V" using superset_gen by auto + with V O V B BU have "V. V O V U" by blast + moreover + from assms(2) B𝔅 BU obtain W where "W𝔅" and "W converse(U)" + using uniformity_base_props(3) by blast + with U have "converse(U) " unfolding Supersets_def + by auto + ultimately have "id(X) U (V. V O V U) converse(U) " + by simp + } thus ?thesis by simp + qed + ultimately show ?thesis unfolding IsUniformity_def by simp qed +text‹The assumption that $X$ is not empty in uniformity_base_is_base› above is neccessary + as the assertion is false if $X$ is empty.› + +lemma uniform_space_empty: assumes "𝔅 {is a uniform base on} " + shows "¬(Supersets(×,𝔅) {is a uniformity on} )" +proof - + { let = "Supersets(×,𝔅)" + assume " {is a uniformity on} " + from assms have "𝔅={}" using uniformity_base_props(5,6) by force + with {is a uniformity on} have False + using supersets_in_empty unif_filter unfolding IsFilter_def by auto + } thus ?thesis by auto +qed + end diff --git a/docs/IsarMathLib/UniformSpace_ZF_2.html b/docs/IsarMathLib/UniformSpace_ZF_2.html index 2f576c7..d8562d9 100644 --- a/docs/IsarMathLib/UniformSpace_ZF_2.html +++ b/docs/IsarMathLib/UniformSpace_ZF_2.html @@ -46,8 +46,8 @@

Theory UniformSpace_ZF_2

begin text‹ The UniformSpace_ZF› theory defines uniform spaces based on entourages (also called surroundings - sometimes). In this theory we consider an alternative definition based of the - notion of uniform covers. + sometimes). In this theory we consider alternative definitions based of the + notion of uniform covers and pseudometrics. › subsection‹ Uniform covers › @@ -67,9 +67,9 @@

Theory UniformSpace_ZF_2

text‹A cover of a nonempty set must have a nonempty member.› -lemma cover_nonempty: assumes "X0" "P Covers(X)" +lemma cover_nonempty: assumes "X0" "P Covers(X)" shows "UP. U0" - using assms unfolding Covers_def by blast + using assms unfolding Covers_def by blast text‹ A "star" of $R$ with respect to $\mathcal{R}$ is the union of all $S\in \mathcal{R}$ that intersect $R$. › @@ -79,8 +79,8 @@

Theory UniformSpace_ZF_2

text‹An element of $\mathcal{R}$ is a subset of its star with respect to $\mathcal{R}$. › -lemma element_subset_star: assumes "UP" shows "U Star(U,P)" - using assms unfolding Star_def by auto +lemma element_subset_star: assumes "UP" shows "U Star(U,P)" + using assms unfolding Star_def by auto text‹An alternative formula for star of a singleton.› @@ -89,13 +89,13 @@

Theory UniformSpace_ZF_2

text‹Star of a larger set is larger.› -lemma star_mono: assumes "UV" shows "Star(U,P) Star(V,P)" - using assms unfolding Star_def by blast +lemma star_mono: assumes "UV" shows "Star(U,P) Star(V,P)" + using assms unfolding Star_def by blast text‹In particular, star of a set is larger than star of any singleton in that set.› -corollary star_single_mono: assumes "xU" shows "Star({x},P) Star(U,P)" - using assms star_mono by auto +corollary star_single_mono: assumes "xU" shows "Star({x},P) Star(U,P)" + using assms star_mono by auto text‹A cover $\mathcal{R}$ (of $X$) is said to be a "barycentric refinement" of a cover $\mathcal{C}$ iff for every $x\in X$ the star of $\{x\}$ in $\mathcal{R}$ is contained @@ -108,9 +108,9 @@

Theory UniformSpace_ZF_2

text‹A cover is a barycentric refinement of the collection of stars of the singletons $\{x \}$ as $x$ ranges over $X$.› -lemma singl_star_bary: +lemma singl_star_bary: assumes "P Covers(X)" shows "P <B {Star({x},P). xX}" - using assms unfolding Covers_def IsBarycentricRefinement_def by blast + using assms unfolding Covers_def IsBarycentricRefinement_def by blast text‹ A cover $\mathcal{R}$ is a "star refinement" of a cover $\mathcal{C}$ iff for each $R\in \mathcal{R}$ there is a $C\in \mathcal{C}$ such that the star of $R$ with @@ -122,18 +122,18 @@

Theory UniformSpace_ZF_2

text‹Every cover star-refines the trivial cover $\{ X\}$. › -lemma cover_stref_triv: assumes "P Covers(X)" shows "P <* {X}" - using assms unfolding Star_def IsStarRefinement_def Covers_def by auto +lemma cover_stref_triv: assumes "P Covers(X)" shows "P <* {X}" + using assms unfolding Star_def IsStarRefinement_def Covers_def by auto text‹Star refinement implies barycentric refinement. › -lemma star_is_bary: assumes "QCovers(X)" and "Q <* P" +lemma star_is_bary: assumes "QCovers(X)" and "Q <* P" shows "Q <B P" proof - - from assms(1) have "Q = X" unfolding Covers_def by simp + from assms(1) have "Q = X" unfolding Covers_def by simp { fix x assume "xX" with Q = X obtain R where "RQ" and "xR" by auto - with assms(2) obtain U where "UP" and "Star(R,Q) U" + with assms(2) obtain U where "UP" and "Star(R,Q) U" unfolding IsStarRefinement_def by auto from xR Star(R,Q) U have "Star({x},Q) U" using star_single_mono by blast @@ -144,27 +144,27 @@

Theory UniformSpace_ZF_2

text‹ Barycentric refinement of a barycentric refinement is a star refinement. › -lemma bary_bary_star: +lemma bary_bary_star: assumes "PCovers(X)" "QCovers(X)" "RCovers(X)" "P <B Q" "Q <B R" "X0" shows "P <* R" proof - { fix U assume "UP" { assume "U = 0" then have "Star(U,P) = 0" unfolding Star_def by simp - from assms(6,3) obtain V where "VR" using cover_nonempty by auto + from assms(6,3) obtain V where "VR" using cover_nonempty by auto with Star(U,P) = 0 have "VR. Star(U,P) V" by auto } moreover { assume "U0" then obtain x0 where "x0U" by auto - with assms(1,2,5) UP obtain V where "VR" and "Star({x0},Q) V" + with assms(1,2,5) UP obtain V where "VR" and "Star({x0},Q) V" unfolding Covers_def IsBarycentricRefinement_def by auto have "Star(U,P) V" proof - { fix W assume "WP" and "WU 0" from WU 0 obtain x where "xWU" by auto - with assms(2) UP have "xP" by auto - with assms(4) obtain C where "CQ" and "Star({x},P) C" + with assms(2) UP have "xP" by auto + with assms(4) obtain C where "CQ" and "Star({x},P) C" unfolding IsBarycentricRefinement_def by blast with UP WP xWU x0U Star({x0},Q) V have "WV" unfolding Star_def by blast @@ -199,41 +199,41 @@

Theory UniformSpace_ZF_2

text‹A family of uniform covers contain the trivial cover $\{ X\}$.› -lemma unicov_contains_triv: assumes "Θ {are uniform covers of} X" +lemma unicov_contains_triv: assumes "Θ {are uniform covers of} X" shows "{X} Θ" proof - - from assms obtain where "Θ" unfolding AreUniformCovers_def by blast - with assms show ?thesis using cover_stref_triv + from assms obtain where "Θ" unfolding AreUniformCovers_def by blast + with assms show ?thesis using cover_stref_triv unfolding AreUniformCovers_def Covers_def by auto qed text‹If $\Theta$ are uniform covers of $X$ then we can recover $X$ from $\Theta$ by taking $\bigcup\bigcup \Theta$. › -lemma space_from_unicov: assumes "Θ {are uniform covers of} X" shows "X = Θ" +lemma space_from_unicov: assumes "Θ {are uniform covers of} X" shows "X = Θ" proof - from assms show "X Θ" using unicov_contains_triv + from assms show "X Θ" using unicov_contains_triv unfolding AreUniformCovers_def by auto - from assms show "Θ X" unfolding AreUniformCovers_def Covers_def + from assms show "Θ X" unfolding AreUniformCovers_def Covers_def by auto qed text‹ Every uniform cover has a star refinement. › -lemma unicov_has_star_ref: +lemma unicov_has_star_ref: assumes "Θ {are uniform covers of} X" and "PΘ" shows "QΘ. (Q <* P)" - using assms unfolding AreUniformCovers_def by blast + using assms unfolding AreUniformCovers_def by blast text‹ In particular, every uniform cover has a barycentric refinement. › -corollary unicov_has_bar_ref: +corollary unicov_has_bar_ref: assumes "Θ {are uniform covers of} X" and "PΘ" shows "QΘ. (Q <B P)" proof - - from assms obtain Q where "QΘ" and "Q <* P" + from assms obtain Q where "QΘ" and "Q <* P" using unicov_has_star_ref by blast - with assms show ?thesis + with assms show ?thesis unfolding AreUniformCovers_def using star_is_bary by blast qed @@ -242,27 +242,27 @@

Theory UniformSpace_ZF_2

shows that in order for $Q$ to be a uniform cover it is sufficient that $P$ is a barycentric refinement of $Q$. ›
-lemma unicov_bary_cov: +lemma unicov_bary_cov: assumes "Θ {are uniform covers of} X" "PΘ" "Q Covers(X)" "P <B Q" and "X0" shows "QΘ" proof - - from assms(1,2) obtain R where "RΘ" and "R <B P" + from assms(1,2) obtain R where "RΘ" and "R <B P" using unicov_has_bar_ref by blast - from assms(1,2,3) RΘ have + from assms(1,2,3) RΘ have "P Covers(X)" "Q Covers(X)" "R Covers(X)" unfolding AreUniformCovers_def by auto - with assms(1,3,4,5) RΘ R <B P show ?thesis + with assms(1,3,4,5) RΘ R <B P show ?thesis using bary_bary_star unfolding AreUniformCovers_def by auto qed text‹ A technical lemma to simplify proof of the uniformity_from_unicov› theorem. › -lemma star_ref_mem: assumes "UP" "P<*Q" and "{W×W. WQ} A" +lemma star_ref_mem: assumes "UP" "P<*Q" and "{W×W. WQ} A" shows "U×U A" proof - - from assms(1,2) obtain W where "WQ" and "{SP. SU 0} W" + from assms(1,2) obtain W where "WQ" and "{SP. SU 0} W" unfolding IsStarRefinement_def Star_def by auto - with assms(1,3) show "U×U A" by blast + with assms(1,3) show "U×U A" by blast qed text‹An identity related to square (in the sense of composition) of a relation of the @@ -283,29 +283,29 @@

Theory UniformSpace_ZF_2

text‹A somewhat technical identity about the square of a symmetric relation: › -lemma rel_sq_image: +lemma rel_sq_image: assumes "W = converse(W)" "domain(W) X" shows "Star({x},{W``{t}. tX}) = (W O W)``{x}" proof - have I: "Star({x},{W``{t}. tX}) = {S{W``{t}. tX}. xS}" + have I: "Star({x},{W``{t}. tX}) = {S{W``{t}. tX}. xS}" unfolding Star_def by auto { fix y assume "y Star({x},{W``{t}. tX})" - with I obtain S where "yS" "xS" "S {W``{t}. tX}" by auto + with I obtain S where "yS" "xS" "S {W``{t}. tX}" by auto from S {W``{t}. tX} obtain t where "tX" and "S = W``{t}" by auto with xS yS have "t,x W" and "t,y W" by auto from t,x W have "x,t converse(W)" by auto - with assms(1) t,y W have "y (W O W)``{x}" + with assms(1) t,y W have "y (W O W)``{x}" using rel_compdef by auto } then show "Star({x},{W``{t}. tX}) (W O W)``{x}" by blast { fix y assume "y(W O W)``{x}" then obtain t where "x,t W" and "t,y W" using rel_compdef by auto - from assms(2) t,y W have "tX" by auto + from assms(2) t,y W have "tX" by auto from x,t W have "t,x converse(W)" by auto - with assms(1) I t,y W tX have "y Star({x},{W``{t}. tX})" + with assms(1) I t,y W tX have "y Star({x},{W``{t}. tX})" by auto } then show "(W O W)``{x} Star({x},{W``{t}. tX})" by blast @@ -321,15 +321,15 @@

Theory UniformSpace_ZF_2

text‹For any member $P$ of a cover $\Theta$ the set $\bigcup \{U\times U : U\in P\}$ is a member of UniformityFromUniCov(X,Θ)›. › -lemma basic_unif: assumes "Θ Covers(X)" "PΘ" +lemma basic_unif: assumes "Θ Covers(X)" "PΘ" shows "{U×U. UP} UniformityFromUniCov(X,Θ)" - using assms unfolding UniformityFromUniCov_def Supersets_def Covers_def + using assms unfolding UniformityFromUniCov_def Supersets_def Covers_def by blast text‹If $\Theta$ is a family of uniform covers of $X$ then UniformityFromUniCov(X,Θ)› is a uniformity on $X$ › -theorem uniformity_from_unicov: +theorem uniformity_from_unicov: assumes "Θ {are uniform covers of} X" "X0" shows "UniformityFromUniCov(X,Θ) {is a uniformity on} X" proof - @@ -342,13 +342,13 @@

Theory UniformSpace_ZF_2

then obtain P where "PΘ" and "0 = {U×U. UP}" unfolding UniformityFromUniCov_def Supersets_def by auto hence "P = 0" by auto - with assms PΘ have False unfolding AreUniformCovers_def Covers_def + with assms PΘ have False unfolding AreUniformCovers_def Covers_def by auto } thus ?thesis by auto qed moreover have "X×X " proof - - from assms have "X×X {{U×U. UP}. PΘ}" + from assms have "X×X {{U×U. UP}. PΘ}" using unicov_contains_triv unfolding AreUniformCovers_def by auto then show ?thesis unfolding Supersets_def UniformityFromUniCov_def @@ -362,15 +362,15 @@

Theory UniformSpace_ZF_2

then have "AB X×X" unfolding UniformityFromUniCov_def Supersets_def by auto from A B obtain PA PB where - "PAΘ" "PBΘ" and I:"{U×U. UPA} A" "{U×U. UPB} B" + "PAΘ" "PBΘ" and I:"{U×U. UPA} A" "{U×U. UPB} B" unfolding UniformityFromUniCov_def Supersets_def by auto - from assms(1) PAΘ PBΘ obtain P + from assms(1) PAΘ PBΘ obtain P where "PΘ" and "P<*PA" and "P<*PB" unfolding AreUniformCovers_def by blast have "{U×U. UP} AB" proof - { fix U assume "UP" - with P<*PA P<*PB I have "U×U A" and "U×U B" + with P<*PA P<*PB I have "U×U A" and "U×U B" using star_ref_mem by auto } thus ?thesis by blast qed @@ -397,24 +397,24 @@

Theory UniformSpace_ZF_2

unfolding UniformityFromUniCov_def Supersets_def by auto have "id(X)A" proof - - from assms(1) PΘ have "P = X" unfolding AreUniformCovers_def Covers_def + from assms(1) PΘ have "P = X" unfolding AreUniformCovers_def Covers_def by auto with {U×U. UP} A show ?thesis by auto qed moreover have "B. B O B A" proof - - from assms(1) PΘ have "{U×U. UP} " + from assms(1) PΘ have "{U×U. UP} " unfolding AreUniformCovers_def Covers_def UniformityFromUniCov_def Supersets_def by auto - from assms(1) PΘ obtain Q where "QΘ" and "Q <* P" using unicov_has_star_ref + from assms(1) PΘ obtain Q where "QΘ" and "Q <* P" using unicov_has_star_ref by blast let ?B = "{U×U. UQ}" - from assms(1) QΘ have "?B " + from assms(1) QΘ have "?B " unfolding AreUniformCovers_def Covers_def UniformityFromUniCov_def Supersets_def by auto moreover have "?B O ?B A" proof - - have II: "?B O ?B = {U×Star(U,Q). UQ}" using rel_square_starr + have II: "?B O ?B = {U×Star(U,Q). UQ}" using rel_square_starr by simp have "UQ. VP. U×Star(U,Q) V×V" proof @@ -427,7 +427,7 @@

Theory UniformSpace_ZF_2

qed hence "{U×Star(U,Q). UQ} {V×V. VP}" by blast with {V×V. VP} A have "{U×Star(U,Q). UQ} A" by blast - with II show ?thesis by simp + with II show ?thesis by simp qed ultimately show ?thesis by auto qed @@ -453,7 +453,7 @@

Theory UniformSpace_ZF_2

of
UniCovFromUniformity› we get an alternative definition of the operation that creates a family of uniform covers from a uniformity. Just a curiosity, not used anywhere.›
-lemma UniCovFromUniformityDef: assumes "X0" +lemma UniCovFromUniformityDef: assumes "X0" shows "UniCovFromUniformity(X,Φ) = (UΦ.xX. {PCovers(X). AP. U``({x}) A})" proof - have "{PCovers(X). UΦ.xX.AP. U``({x}) A} = @@ -462,7 +462,7 @@

Theory UniformSpace_ZF_2

{ fix P assume "P{PCovers(X). UΦ.xX.AP. U``({x}) A}" then have "PCovers(X)" and "UΦ.xX.AP. U``({x}) A" by auto then obtain U where "UΦ" and "xX.AP. U``({x}) A" by auto - with assms PCovers(X) have "P (xX. {PCovers(X). AP. U``({x}) A})" + with assms PCovers(X) have "P (xX. {PCovers(X). AP. U``({x}) A})" by auto with UΦ have "P(UΦ.xX. {PCovers(X). AP. U``({x}) A})" by blast @@ -473,7 +473,7 @@

Theory UniformSpace_ZF_2

{ fix P assume "P(UΦ.xX. {PCovers(X). AP. U``({x}) A})" then obtain U where "UΦ" "P (xX. {PCovers(X). AP. U``({x}) A})" by auto - with assms have "PCovers(X)" and "xX.AP. U``({x}) A" by auto + with assms have "PCovers(X)" and "xX.AP. U``({x}) A" by auto with UΦ have "P{PCovers(X). UΦ.xX.AP. U``({x}) A}" by auto } then show "(UΦ.xX. {PCovers(X). AP. U``({x}) A}) @@ -485,23 +485,23 @@

Theory UniformSpace_ZF_2

text‹If $\Phi$ is a (diagonal) uniformity on $X$, then covers of the form $\{ W\{ x\} : x\in X\}$ are members of UniCovFromUniformity(X,Φ)›. › -lemma cover_image: +lemma cover_image: assumes "Φ {is a uniformity on} X" "WΦ" shows "{W``{x}. xX} UniCovFromUniformity(X,Φ)" proof - let ?P = "{W``{x}. xX}" have "?P Covers(X)" proof - - from assms have "W X×X" and "?P Pow(Pow(X))" + from assms have "W X×X" and "?P Pow(Pow(X))" using entourage_props(1) by auto moreover have "?P = X" proof from W X×X show "?P X" by auto - from assms show "X ?P" using neigh_not_empty(2) by auto + from assms show "X ?P" using neigh_not_empty(2) by auto qed ultimately show ?thesis unfolding Covers_def by simp qed - moreover from assms(2) have "WΦ. xX. A?P. W``{x} A" + moreover from assms(2) have "WΦ. xX. A?P. W``{x} A" by auto ultimately show ?thesis unfolding UniCovFromUniformity_def by simp @@ -510,39 +510,39 @@

Theory UniformSpace_ZF_2

text‹If $\Phi$ is a (diagonal) uniformity on $X$, then every two elements of UniCovFromUniformity(X,Φ)› have a common barycentric refinement.› -lemma common_bar_refinemnt: +lemma common_bar_refinemnt: assumes "Φ {is a uniformity on} X" "Θ = UniCovFromUniformity(X,Φ)" "𝒞Θ" "𝒟Θ" shows "Θ.( <B 𝒞) ( <B 𝒟)" proof - - from assms(2,3) obtain U where "UΦ" and I: "xX.C𝒞. U``{x} C" + from assms(2,3) obtain U where "UΦ" and I: "xX.C𝒞. U``{x} C" unfolding UniCovFromUniformity_def by auto - from assms(2,4) obtain V where "VΦ" and II: "xX.D𝒟. V``{x} D" + from assms(2,4) obtain V where "VΦ" and II: "xX.D𝒟. V``{x} D" unfolding UniCovFromUniformity_def by auto - from assms(1) UΦ VΦ have "UV Φ" + from assms(1) UΦ VΦ have "UV Φ" unfolding IsUniformity_def IsFilter_def by auto - with assms(1) obtain W where "WΦ" and "W O W UV" and "W=converse(W)" + with assms(1) obtain W where "WΦ" and "W O W UV" and "W=converse(W)" using half_size_symm by blast - from assms(1) WΦ have "domain(W) X" + from assms(1) WΦ have "domain(W) X" unfolding IsUniformity_def IsFilter_def by auto let ?P = "{W``{t}. tX}" have "?PΘ" "?P <B 𝒞" "?P <B 𝒟" proof - - from assms(1,2) WΦ show "?PΘ" using cover_image by simp - with assms(2) have "?P = X" unfolding UniCovFromUniformity_def Covers_def + from assms(1,2) WΦ show "?PΘ" using cover_image by simp + with assms(2) have "?P = X" unfolding UniCovFromUniformity_def Covers_def by simp { fix x assume "xX" from W=converse(W) domain(W) X W O W UV have "Star({x},?P) U``{x}" and "Star({x},?P) V``{x}" using rel_sq_image by auto - from xX I obtain C where "C𝒞" and "U``{x} C" + from xX I obtain C where "C𝒞" and "U``{x} C" by auto with Star({x},?P) U``{x} C𝒞 have "C𝒞. Star({x},?P) C" by auto moreover - from xX II obtain D where "D𝒟" and "V``{x} D" + from xX II obtain D where "D𝒟" and "V``{x} D" by auto with Star({x},?P) V``{x} D𝒟 have "D𝒟. Star({x},?P) D" by auto @@ -559,23 +559,23 @@

Theory UniformSpace_ZF_2

text‹If $\Phi$ is a (diagonal) uniformity on $X$, then every element of UniCovFromUniformity(X,Φ)› has a barycentric refinement there.› -corollary bar_refinement_ex: +corollary bar_refinement_ex: assumes "Φ {is a uniformity on} X" "Θ = UniCovFromUniformity(X,Φ)" "𝒞 Θ" shows "Θ. ( <B 𝒞)" - using assms common_bar_refinemnt by blast + using assms common_bar_refinemnt by blast text‹ If $\Phi$ is a (diagonal) uniformity on $X$, then UniCovFromUniformity(X,Φ)› is a family of uniform covers.› -theorem unicov_from_uniformity: assumes "Φ {is a uniformity on} X" and "X0" +theorem unicov_from_uniformity: assumes "Φ {is a uniformity on} X" and "X0" shows "UniCovFromUniformity(X,Φ) {are uniform covers of} X" proof - let = "UniCovFromUniformity(X,Φ)" - from assms(1) have " Covers(X)" unfolding UniCovFromUniformity_def + from assms(1) have " Covers(X)" unfolding UniCovFromUniformity_def by auto moreover - from assms(1) have "{X} " + from assms(1) have "{X} " unfolding Covers_def IsUniformity_def IsFilter_def UniCovFromUniformity_def by auto hence " 0" by auto @@ -584,10 +584,10 @@

Theory UniformSpace_ZF_2

{ fix 𝒞 assume "" "𝒞Covers(X)" " <* 𝒞" have "𝒞" proof - - from obtain U where "UΦ" and I: "xX.R. U``({x}) R" + from obtain U where "UΦ" and I: "xX.R. U``({x}) R" unfolding UniCovFromUniformity_def by auto { fix x assume "xX" - with I obtain R where "R" and "U``({x}) R" by auto + with I obtain R where "R" and "U``({x}) R" by auto from R <* 𝒞 obtain C where "C𝒞" and "Star(R,) C" unfolding IsStarRefinement_def by auto with U``({x}) R R have "U``({x}) C" @@ -602,14 +602,14 @@

Theory UniformSpace_ZF_2

moreover have "𝒞.𝒟..( <* 𝒞) ( <* 𝒟)" proof - { fix 𝒞 𝒟 assume "𝒞" "𝒟" - with assms(1) obtain P where "P" and "P <B 𝒞" "P <B 𝒟" + with assms(1) obtain P where "P" and "P <B 𝒞" "P <B 𝒟" using common_bar_refinemnt by blast - from assms(1) P obtain where "" and " <B P" + from assms(1) P obtain where "" and " <B P" using bar_refinement_ex by blast from P 𝒞 𝒟 have "P Covers(X)" " Covers(X)" "𝒞 Covers(X)" "𝒟 Covers(X)" unfolding UniCovFromUniformity_def by auto - with assms(2) <B P P <B 𝒞 P <B 𝒟 have " <* 𝒞" and " <* 𝒟" + with assms(2) <B P P <B 𝒞 P <B 𝒟 have " <* 𝒞" and " <* 𝒟" using bary_bary_star by auto with have ".( <* 𝒞) ( <* 𝒟)" by auto } thus ?thesis by simp @@ -619,19 +619,19 @@

Theory UniformSpace_ZF_2

text‹ The UniCovFromUniformity› operation is the inverse of UniformityFromUniCov›. › -theorem unicov_from_unif_inv: assumes "Θ {are uniform covers of} X" "X0" +theorem unicov_from_unif_inv: assumes "Θ {are uniform covers of} X" "X0" shows "UniCovFromUniformity(X,UniformityFromUniCov(X,Θ)) = Θ" proof let = "UniformityFromUniCov(X,Θ)" let ?L = "UniCovFromUniformity(X,)" - from assms have I: " {is a uniformity on} X" + from assms have I: " {is a uniformity on} X" using uniformity_from_unicov by simp - with assms(2) have II: "?L {are uniform covers of} X" + with assms(2) have II: "?L {are uniform covers of} X" using unicov_from_uniformity by simp { fix P assume "P?L" - with I obtain Q where "Q?L" and "Q <B P" + with I obtain Q where "Q?L" and "Q <B P" using bar_refinement_ex by blast - from Q?L obtain U where "U" and III:"xX.AQ. U``{x} A" + from Q?L obtain U where "U" and III:"xX.AQ. U``{x} A" unfolding UniCovFromUniformity_def by auto from U have "U Supersets(X×X,{{U×U. UP}. PΘ})" unfolding UniformityFromUniCov_def by simp @@ -640,13 +640,13 @@

Theory UniformSpace_ZF_2

then obtain C where "C {{U×U. UP}. PΘ}" and "CB" by auto then obtain R where "RΘ" and "C = {V×V. VR}" by auto with CB BU have "{V×V. VR} U" by auto - from assms(1) II P?L Q?L RΘ have - IV: "PCovers(X)" "QCovers(X)" "RCovers(X)" + from assms(1) II P?L Q?L RΘ have + IV: "PCovers(X)" "QCovers(X)" "RCovers(X)" unfolding AreUniformCovers_def by auto have "R <B Q" proof - { fix x assume "xX" - with III obtain A where "AQ" and "U``{x} A" by auto + with III obtain A where "AQ" and "U``{x} A" by auto with {V×V. VR} U have "({V×V. VR})``{x} A" by auto with AQ have "AQ. Star({x},R) A" using star_singleton by auto @@ -656,38 +656,38 @@

Theory UniformSpace_ZF_2

ultimately show ?thesis unfolding IsBarycentricRefinement_def by simp qed - with assms(2) Q <B P IV have "R <* P" using bary_bary_star by simp - with assms(1) RΘ PCovers(X) have "PΘ" + with assms(2) Q <B P IV have "R <* P" using bary_bary_star by simp + with assms(1) RΘ PCovers(X) have "PΘ" unfolding AreUniformCovers_def by simp } thus "?LΘ" by auto { fix P assume "PΘ" - with assms(1) have "P Covers(X)" + with assms(1) have "P Covers(X)" unfolding AreUniformCovers_def by auto - from assms(1) PΘ obtain Q where "Q Θ" and "Q <B P" + from assms(1) PΘ obtain Q where "Q Θ" and "Q <B P" using unicov_has_bar_ref by blast let ?A = "{V×V. VQ}" have "?A " proof - - from assms(1) QΘ have "?A X×X" and "?A {{V×V. VQ}. QΘ}" + from assms(1) QΘ have "?A X×X" and "?A {{V×V. VQ}. QΘ}" unfolding AreUniformCovers_def Covers_def by auto then show ?thesis using superset_gen unfolding UniformityFromUniCov_def by auto qed - with I obtain B where "B" "B O B ?A" and "B=converse(B)" + with I obtain B where "B" "B O B ?A" and "B=converse(B)" using half_size_symm by blast let ?R = "{B``{x}. xX}" - from I II B have "?R?L" and "?R =X" + from I II B have "?R?L" and "?R =X" using cover_image unfolding UniCovFromUniformity_def Covers_def by auto have "?R <B P" proof - { fix x assume "xX" - from assms(1) Q Θ have "Q = X" + from assms(1) Q Θ have "Q = X" unfolding AreUniformCovers_def Covers_def by auto with Q <B P xX obtain C where "CP" and "Star({x},Q) C" unfolding IsBarycentricRefinement_def by auto - from B=converse(B) I B have "Star({x},?R) = (B O B)``{x}" + from B=converse(B) I B have "Star({x},?R) = (B O B)``{x}" using uni_domain rel_sq_image by auto moreover from (B O B) ?A have "(B O B)``{x} ?A``{x}" by blast moreover have "?A``{x} = Star({x},Q)" using star_singleton by simp @@ -697,59 +697,59 @@

Theory UniformSpace_ZF_2

} with ?R = X show ?thesis unfolding IsBarycentricRefinement_def by auto qed - with assms(2) II P Covers(X) ?R?L ?R <B P have "P?L" + with assms(2) II P Covers(X) ?R?L ?R <B P have "P?L" using unicov_bary_cov by simp } thus "Θ?L" by auto qed text‹The UniformityFromUniCov› operation is the inverse of UniCovFromUniformity›. › -theorem unif_from_unicov_inv: assumes "Φ {is a uniformity on} X" "X0" +theorem unif_from_unicov_inv: assumes "Φ {is a uniformity on} X" "X0" shows "UniformityFromUniCov(X,UniCovFromUniformity(X,Φ)) = Φ" proof let = "UniCovFromUniformity(X,Φ)" let ?L = "UniformityFromUniCov(X,)" - from assms have I: " {are uniform covers of} X" + from assms have I: " {are uniform covers of} X" using unicov_from_uniformity by simp - with assms have II: "?L {is a uniformity on} X" + with assms have II: "?L {is a uniformity on} X" using uniformity_from_unicov by simp { fix A assume "AΦ" - with assms(1) obtain B where "BΦ" "B O B A" and "B = converse(B)" + with assms(1) obtain B where "BΦ" "B O B A" and "B = converse(B)" using half_size_symm by blast - from assms(1) AΦ have "A X×X" using uni_domain(1) + from assms(1) AΦ have "A X×X" using uni_domain(1) by simp let ?P = "{B``{x}. xX}" - from assms(1) BΦ have "?P" using cover_image + from assms(1) BΦ have "?P" using cover_image by simp let ?C = "{U×U. U?P}" - from I ?P have "?C?L" + from I ?P have "?C?L" unfolding AreUniformCovers_def using basic_unif by blast - from assms(1) BΦ B = converse(B) B O B A have "?C A" + from assms(1) BΦ B = converse(B) B O B A have "?C A" using uni_domain(2) symm_sq_prod_image by simp - with II A X×X ?C?L have "A?L" + with II A X×X ?C?L have "A?L" unfolding IsUniformity_def IsFilter_def by simp } thus "Φ?L" by auto { fix A assume "A?L" - with II have "A X×X" using entourage_props(1) by simp + with II have "A X×X" using entourage_props(1) by simp from A?L obtain P where "P" and "{U×U. UP} A" unfolding UniformityFromUniCov_def Supersets_def by blast - from P obtain B where "BΦ" and III: "xX. VP. B``{x} V" + from P obtain B where "BΦ" and III: "xX. VP. B``{x} V" unfolding UniCovFromUniformity_def by auto have "BA" proof - - from assms(1) BΦ have "B {B``{x}×B``{x}. xX}" + from assms(1) BΦ have "B {B``{x}×B``{x}. xX}" using entourage_props(1,2) refl_union_singl_image by simp moreover have "{B``{x}×B``{x}. xX} A" proof - { fix x assume "xX" - with III obtain V where "VP" and "B``{x} V" by auto + with III obtain V where "VP" and "B``{x} V" by auto hence "B``{x}×B``{x} {U×U. UP}" by auto } hence "{B``{x}×B``{x}. xX} {U×U. UP}" by blast with {U×U. UP} A show ?thesis by blast qed ultimately show ?thesis by auto qed - with assms(1) BΦ A X×X have "AΦ" + with assms(1) BΦ A X×X have "AΦ" unfolding IsUniformity_def IsFilter_def by simp } thus "?LΦ" by auto qed diff --git a/docs/IsarMathLib/ZF1.html b/docs/IsarMathLib/ZF1.html index a7d7c7d..7812f55 100644 --- a/docs/IsarMathLib/ZF1.html +++ b/docs/IsarMathLib/ZF1.html @@ -69,42 +69,48 @@

Theory ZF1

with assms show "x B" by auto qed +text‹In ZF set theory the zero of natural numbers is the same as the empty set. + In the next abbreviation we declare that we want $0$ and $\emptyset$ to be synonyms + so that we can use $\emptyset$ instead of $0$ when appropriate. › + +abbreviation empty_set ("") where " 0" + text‹If all sets of a nonempty collection are the same, then its union is the same.› -lemma ZF1_1_L1: assumes "C0" and "yC. b(y) = A" - shows "(yC. b(y)) = A" using assms by blast +lemma ZF1_1_L1: assumes "C" and "yC. b(y) = A" + shows "(yC. b(y)) = A" using assms by blast text‹The union af all values of a constant meta-function belongs to the same set as the constant.› -lemma ZF1_1_L2: assumes A1:"C0" and A2: "xC. b(x) A" - and A3: "x y. xC yC b(x) = b(y)" +lemma ZF1_1_L2: assumes A1:"C" and A2: "xC. b(x) A" + and A3: "x y. xC yC b(x) = b(y)" shows "(xC. b(x))A" proof - - from A1 obtain x where D1: "xC" by auto - with A3 have "yC. b(y) = b(x)" by blast - with A1 have "(yC. b(y)) = b(x)" + from A1 obtain x where D1: "xC" by auto + with A3 have "yC. b(y) = b(x)" by blast + with A1 have "(yC. b(y)) = b(x)" using ZF1_1_L1 by simp - with D1 A2 show ?thesis by simp + with D1 A2 show ?thesis by simp qed text‹If two meta-functions are the same on a cartesian product, then the subsets defined by them are the same. I am surprised Isabelle can not handle this automatically.› -lemma ZF1_1_L4: assumes A1: "xX.yY. a(x,y) = b(x,y)" +lemma ZF1_1_L4: assumes A1: "xX.yY. a(x,y) = b(x,y)" shows "{a(x,y). x,y X×Y} = {b(x,y). x,y X×Y}" proof show "{a(x, y). x,y X × Y} {b(x, y). x,y X × Y}" proof fix z assume "z {a(x, y) . x,y X × Y}" - with A1 show "z {b(x,y).x,y X×Y}" by auto + with A1 show "z {b(x,y).x,y X×Y}" by auto qed show "{b(x, y). x,y X × Y} {a(x, y). x,y X × Y}" proof fix z assume "z {b(x, y). x,y X × Y}" - with A1 show "z {a(x,y).x,y X×Y}" by auto + with A1 show "z {a(x,y).x,y X×Y}" by auto qed qed @@ -114,19 +120,19 @@

Theory ZF1

the set definition varies over
p∈X×Y› rather than ⟨ x,y⟩∈X×Y›.›
-lemma ZF1_1_L4A: assumes A1: "xX.yY. a( x,y) = b(x,y)" +lemma ZF1_1_L4A: assumes A1: "xX.yY. a( x,y) = b(x,y)" shows "{a(p). p X×Y} = {b(x,y). x,y X×Y}" proof { fix z assume "z {a(p). pX×Y}" - then obtain p where D1: "z=a(p)" "pX×Y" by auto + then obtain p where D1: "z=a(p)" "pX×Y" by auto let ?x = "fst(p)" let ?y = "snd(p)" - from A1 D1 have "z {b(x,y). x,y X×Y}" by auto + from A1 D1 have "z {b(x,y). x,y X×Y}" by auto } then show "{a(p). p X×Y} {b(x,y). x,y X×Y}" by blast next { fix z assume "z {b(x,y). x,y X×Y}" - then obtain x y where D1: "x,y X×Y" "z=b(x,y)" by auto + then obtain x y where D1: "x,y X×Y" "z=b(x,y)" by auto let ?p = " x,y" - from A1 D1 have "?pX×Y" "z = a(?p)" by auto + from A1 D1 have "?pX×Y" "z = a(?p)" by auto then have "z {a(p). p X×Y}" by auto } then show "{b(x,y). x,y X×Y} {a(p). p X×Y}" by blast qed @@ -134,43 +140,50 @@

Theory ZF1

text‹A lemma about inclusion in cartesian products. Included here to remember that we need the $U\times V \neq \emptyset$ assumption.› -lemma prod_subset: assumes "U×V0" "U×V X×Y" shows "UX" and "VY" - using assms by auto +lemma prod_subset: assumes "U×V" "U×V X×Y" shows "UX" and "VY" + using assms by auto text‹A technical lemma about sections in cartesian products.› -lemma section_proj: assumes "A X×Y" and "U×V A" and "x U" "y V" +lemma section_proj: assumes "A X×Y" and "U×V A" and "x U" "y V" shows "U {tX. t,y A}" and "V {tY. x,t A}" - using assms by auto + using assms by auto text‹If two meta-functions are the same on a set, then they define the same set by separation.› -lemma ZF1_1_L4B: assumes "xX. a(x) = b(x)" +lemma ZF1_1_L4B: assumes "xX. a(x) = b(x)" shows "{a(x). xX} = {b(x). xX}" - using assms by simp + using assms by simp text‹A set defined by a constant meta-function is a singleton.› -lemma ZF1_1_L5: assumes "X0" and "xX. b(x) = c" - shows "{b(x). xX} = {c}" using assms by blast +lemma ZF1_1_L5: assumes "X" and "xX. b(x) = c" + shows "{b(x). xX} = {c}" using assms by blast text‹Most of the time, auto› does this job, but there are strange cases when the next lemma is needed.› -lemma subset_with_property: assumes "Y = {xX. b(x)}" +lemma subset_with_property: assumes "Y = {xX. b(x)}" shows "Y X" - using assms by auto + using assms by auto + +text‹If set $A$ is contained in set $B$ and exist elements $x,y$ of the set $A$ + that satisfy a predicate then exist elements of the set $B$ that satisfy the predicate. › + +lemma exist2_subset: assumes "AB" and "xA. yA. φ(x,y)" + shows "xB. yB. φ(x,y)" + using assms by blast text‹We can choose an element from a nonempty set.› -lemma nonempty_has_element: assumes "X0" shows "x. xX" - using assms by auto +lemma nonempty_has_element: assumes "X" shows "x. xX" + using assms by auto (*text{*If after removing an element from a set we get an empty set, then this set must be a singleton.*} -lemma rem_point_empty: assumes "a∈A" and "A-{a} = 0" +lemma rem_point_empty: assumes "a∈A" and "A-{a} = ∅" shows "A = {a}" using assms by auto; *) text‹In Isabelle/ZF the intersection of an empty family is @@ -179,14 +192,14 @@

Theory ZF1

difficult to find. This is one reason we need comments before every theorem: so that we can search for keywords.›
-lemma inter_empty_empty: shows "0 = 0" by (rule Inter_0) +lemma inter_empty_empty: shows " = " by (rule Inter_0) text‹If an intersection of a collection is not empty, then the collection is not empty. We are (ab)using the fact the the intersection of empty collection is defined to be empty.› -lemma inter_nempty_nempty: assumes "A 0" shows "A0" - using assms by auto +lemma inter_nempty_nempty: assumes "A " shows "A" + using assms by auto text‹For two collections $S,T$ of sets we define the product collection as the collections of cartesian products $A\times B$, where $A\in S, B\in T$.› @@ -202,11 +215,11 @@

Theory ZF1

text‹An intersection of subsets is a subset.› -lemma ZF1_1_L7: assumes A1: "I0" and A2: "iI. P(i) X" +lemma ZF1_1_L7: assumes A1: "I" and A2: "iI. P(i) X" shows "( iI. P(i) ) X" proof - - from A1 obtain i0 where "i0 I" by auto - with A2 have "( iI. P(i) ) P(i0)" and "P(i0) X" + from A1 obtain i0 where "i0 I" by auto + with A2 have "( iI. P(i) ) P(i0)" and "P(i0) X" by auto thus "( iI. P(i) ) X" by auto qed @@ -219,17 +232,17 @@

Theory ZF1

text‹Some properties of singletons.› -lemma ZF1_1_L9: assumes A1: "∃! x. xA φ(x)" +lemma ZF1_1_L9: assumes A1: "∃! x. xA φ(x)" shows "a. {xA. φ(x)} = {a}" " {xA. φ(x)} A" "φ( {xA. φ(x)})" proof - - from A1 show "a. {xA. φ(x)} = {a}" by auto - then obtain a where I: "{xA. φ(x)} = {a}" by auto + from A1 show "a. {xA. φ(x)} = {a}" by auto + then obtain a where I: "{xA. φ(x)} = {a}" by auto then have " {xA. φ(x)} = a" by auto moreover - from I have "a {xA. φ(x)}" by simp + from I have "a {xA. φ(x)}" by simp hence "aA" and "φ(a)" by auto ultimately show " {xA. φ(x)} A" and "φ( {xA. φ(x)})" by auto @@ -237,10 +250,10 @@

Theory ZF1

text‹A simple version of ZF1_1_L9›.› -corollary singleton_extract: assumes "∃! x. xA" +corollary singleton_extract: assumes "∃! x. xA" shows "( A) A" proof - - from assms have "∃! x. xA True" by simp + from assms have "∃! x. xA True" by simp then have " {xA. True} A" by (rule ZF1_1_L9) thus "( A) A" by simp qed @@ -248,31 +261,31 @@

Theory ZF1

text‹A criterion for when a set defined by comprehension is a singleton.› lemma singleton_comprehension: - assumes A1: "yX" and A2: "xX. yX. P(x) = P(y)" + assumes A1: "yX" and A2: "xX. yX. P(x) = P(y)" shows "({P(x). xX}) = P(y)" proof - let ?A = "{P(x). xX}" have "∃! c. c ?A" proof - from A1 show "c. c ?A" by auto + from A1 show "c. c ?A" by auto next fix a b assume "a ?A" and "b ?A" then obtain x t where "x X" "a = P(x)" and "t X" "b = P(t)" by auto - with A2 show "a=b" by blast + with A2 show "a=b" by blast qed then have "(?A) ?A" by (rule singleton_extract) then obtain x where "x X" and "(?A) = P(x)" by auto - from A1 A2 x X have "P(x) = P(y)" + from A1 A2 x X have "P(x) = P(y)" by blast with (?A) = P(x) show "(?A) = P(y)" by simp qed text‹Adding an element of a set to that set does not change the set.› -lemma set_elem_add: assumes "xX" shows "X {x} = X" using assms +lemma set_elem_add: assumes "xX" shows "X {x} = X" using assms by auto text‹Here we define a restriction of a collection of sets to a given set. @@ -294,7 +307,7 @@

Theory ZF1

text‹Next we show a technical identity that is used to prove sufficiency of some condition for a collection of sets to be a base for a topology.› -lemma ZF1_1_L10: assumes A1: "UC. AB. U = A" +lemma ZF1_1_L10: assumes A1: "UC. AB. U = A" shows " {{AB. U = A}. UC} = C" proof show "(UC. {A B . U = A}) C" by blast @@ -304,7 +317,7 @@

Theory ZF1

show "x (UC. {A B . U = A})" proof - from x C obtain U where "UC xU" by auto - with A1 obtain A where "AB U = A" by auto + with A1 obtain A where "AB U = A" by auto from UC xU AB U = A show "x (UC. {A B . U = A})" by auto qed @@ -320,54 +333,54 @@

Theory ZF1

text‹If a difference between a set and a singleton is empty, then the set is empty or it is equal to the singleton.› -lemma singl_diff_empty: assumes "A - {x} = 0" - shows "A = 0 A = {x}" - using assms by auto +lemma singl_diff_empty: assumes "A - {x} = " + shows "A = A = {x}" + using assms by auto text‹If a difference between a set and a singleton is the set, then the only element of the singleton is not in the set.› -lemma singl_diff_eq: assumes A1: "A - {x} = A" +lemma singl_diff_eq: assumes A1: "A - {x} = A" shows "x A" proof - have "x A - {x}" by auto - with A1 show "x A" by simp + with A1 show "x A" by simp qed text‹Simple substitution in membership, has to be used by rule in very rare cases.› -lemma eq_mem: assumes "xA" and "y=x" shows "yA" - using assms by simp +lemma eq_mem: assumes "xA" and "y=x" shows "yA" + using assms by simp text‹A basic property of sets defined by comprehension.› -lemma comprehension: assumes "a {xX. p(x)}" - shows "aX" and "p(a)" using assms by auto +lemma comprehension: assumes "a {xX. p(x)}" + shows "aX" and "p(a)" using assms by auto text‹A basic property of a set defined by another type of comprehension.› -lemma comprehension_repl: assumes "y {p(x). xX}" - shows "xX. y = p(x)" using assms by auto +lemma comprehension_repl: assumes "y {p(x). xX}" + shows "xX. y = p(x)" using assms by auto text‹The inverse of the comprehension› lemma.› -lemma mem_cond_in_set: assumes "φ(c)" and "cX" - shows "c {xX. φ(x)}" using assms by blast +lemma mem_cond_in_set: assumes "φ(c)" and "cX" + shows "c {xX. φ(x)}" using assms by blast text‹The image of a set by a greater relation is greater. › -lemma image_rel_mono: assumes "rs" shows "r``(A) s``(A)" - using assms by auto +lemma image_rel_mono: assumes "rs" shows "r``(A) s``(A)" + using assms by auto text‹ A technical lemma about relations: if $x$ is in its image by a relation $U$ and that image is contained in some set $C$, then the image of the singleton $\{ x\}$ by the relation $U \cup C\times C$ equals $C$. › -lemma image_greater_rel: +lemma image_greater_rel: assumes "x U``{x}" and "U``{x} C" shows "(U C×C)``{x} = C" - using assms image_Un_left by blast + using assms image_Un_left by blast text‹Reformulation of the definition of composition of two relations: › @@ -383,7 +396,7 @@

Theory ZF1

text‹An identity for the square (in the sense of composition) of a symmetric relation.› -lemma symm_sq_prod_image: assumes "converse(r) = r" +lemma symm_sq_prod_image: assumes "converse(r) = r" shows "r O r = {(r``{x})×(r``{x}). x domain(r)}" proof { fix p assume "p r O r" @@ -391,7 +404,7 @@

Theory ZF1

with p r O r obtain x where "y,x r" and "x,z r" using rel_compdef by auto from y,x r have "x,y converse(r)" by simp - with assms x,z r y,z = p have "xdomain(r). p (r``{x})×(r``{x})" + with assms x,z r y,z = p have "xdomain(r). p (r``{x})×(r``{x})" by auto } thus "r O r ({(r``{x})×(r``{x}). x domain(r)})" by blast @@ -403,21 +416,28 @@

Theory ZF1

by auto from y r``{x} have "x,y r" by auto then have "y,x converse(r)" by simp - with assms z r``{x} y,z = p have "p r O r" by auto + with assms z r``{x} y,z = p have "p r O r" by auto } thus ?thesis by auto qed } thus "({(r``{x})×(r``{x}). x domain(r)}) r O r" by blast qed +text‹Square of a reflexive relation contains the relation. + Recall that in ZF the identity function on $X$ is the same as the diagonal + of $X\times X$, i.e. $id(X) = \{\langle x,x\rangle : x\in X\}$. › + +lemma refl_square_greater: assumes "r X×X" "id(X) r" + shows "r r O r" using assms by auto + text‹A reflexive relation is contained in the union of products of its singleton images. › -lemma refl_union_singl_image: +lemma refl_union_singl_image: assumes "A X×X" and "id(X)A" shows "A {A``{x}×A``{x}. x X}" proof - { fix p assume "pA" - with assms(1) obtain x y where "xX" "yX" and "p=x,y" by auto - with assms(2) pA have "xX. p A``{x}×A``{x}" by auto + with assms(1) obtain x y where "xX" "yX" and "p=x,y" by auto + with assms(2) pA have "xX. p A``{x}×A``{x}" by auto } thus ?thesis by auto qed @@ -425,21 +445,21 @@

Theory ZF1

symmetric relation $W$ has a nonempty intersection with $R$ then $x$ is in relation $W\circ (R\circ W)$ with $y$. ›
-lemma sym_rel_comp: - assumes "W=converse(W)" and "(W``{x})×(W``{y}) R 0" +lemma sym_rel_comp: + assumes "W=converse(W)" and "(W``{x})×(W``{y}) R " shows "x,y (W O (R O W))" proof - - from assms(2) obtain s t where "sW``{x}" "tW``{y}" and "s,tR" + from assms(2) obtain s t where "sW``{x}" "tW``{y}" and "s,tR" by blast then have "x,s W" and "y,t W" by auto from x,s W s,t R have "x,t R O W" by auto from y,t W have "t,y converse(W)" by blast - with assms(1) x,t R O W show ?thesis by auto + with assms(1) x,t R O W show ?thesis by auto qed text‹ It's hard to believe but there are cases where we have to reference this rule. › -lemma set_mem_eq: assumes "xA" "A=B" shows "xB" using assms by simp +lemma set_mem_eq: assumes "xA" "A=B" shows "xB" using assms by simp text‹Given some family $\mathcal{A}$ of subsets of $X$ we can define the family of supersets of $\mathcal{A}$. › @@ -449,50 +469,71 @@

Theory ZF1

text‹The family itself is in its supersets. › -lemma superset_gen: assumes "AX" "A𝒜" shows "A Supersets(X,𝒜)" - using assms unfolding Supersets_def by auto +lemma superset_gen: assumes "AX" "A𝒜" shows "A Supersets(X,𝒜)" + using assms unfolding Supersets_def by auto + +text‹The whole space is a superset of any nonempty collection of its subsets. › + +lemma space_superset: assumes "𝒜" "𝒜Pow(X)" shows "X Supersets(X,𝒜)" +proof - + from assms(1) obtain A where "A𝒜" by auto + with assms(2) show ?thesis unfolding Supersets_def by auto +qed + +text‹The collection of supersets of an empty set is empty. In particular + the whole space $X$ is not a superset of an empty set. › + +lemma supersets_of_empty: shows "Supersets(X,) = " + unfolding Supersets_def by auto + +text‹However, when the space is empty the collection of supersets does not have + to be empty - the collection of supersets of the singleton collection containing + only the empty set is this collection. › + +lemma supersets_in_empty: shows "Supersets(,{}) = {}" + unfolding Supersets_def by auto text‹This can be done by the auto method, but sometimes takes a long time. › -lemma witness_exists: assumes "xX" and "φ(x)" shows "xX. φ(x)" - using assms by auto +lemma witness_exists: assumes "xX" and "φ(x)" shows "xX. φ(x)" + using assms by auto text‹Another lemma that concludes existence of some set.› -lemma witness_exists1: assumes "xX" "φ(x)" "ψ(x)" +lemma witness_exists1: assumes "xX" "φ(x)" "ψ(x)" shows "xX. φ(x) ψ(x)" - using assms by auto + using assms by auto text‹The next lemma has to be used as a rule in some rare cases. › -lemma exists_in_set: assumes "x. xA φ(x)" shows "xA. φ(x)" - using assms by simp +lemma exists_in_set: assumes "x. xA φ(x)" shows "xA. φ(x)" + using assms by simp text‹If $x$ belongs to a set where a property holds, then the property holds for $x$. This has to be used as rule in rare cases. › -lemma property_holds: assumes "tX. φ(t)" and "xX" - shows "φ(x)" using assms by simp +lemma property_holds: assumes "tX. φ(t)" and "xX" + shows "φ(x)" using assms by simp text‹Set comprehensions defined by equal expressions are the equal. The second assertion is actually about functions, which are sets of pairs as illustrated in lemma fun_is_set_of_pairs› in func1.thy› -lemma set_comp_eq: assumes "xX. p(x) = q(x)" +lemma set_comp_eq: assumes "xX. p(x) = q(x)" shows "{p(x). xX} = {q(x). xX}" and "{x,p(x). xX} = {x,q(x). xX}" - using assms by auto + using assms by auto text‹If every element of a non-empty set $X\subseteq Y$ satisfies a condition then the set of elements of $Y$ that satisfy the condition is non-empty.› -lemma non_empty_cond: assumes "X0" "XY" and "xX. P(x)" - shows "{xY. P(x)} 0" using assms by auto +lemma non_empty_cond: assumes "X" "XY" and "xX. P(x)" + shows "{xY. P(x)} 0" using assms by auto text‹If $z$ is a pair, then the cartesian product of the singletons of its elements is the same as the singleton $\{ z\}$.› -lemma pair_prod: assumes "z = x,y" shows "{x}×{y} = {z}" - using assms by blast +lemma pair_prod: assumes "z = x,y" shows "{x}×{y} = {z}" + using assms by blast text‹In Isabelle/ZF the set difference is written with a minus sign $A-B$ because the standard backslash character is reserved for other purposes. @@ -501,11 +542,6 @@

Theory ZF1

abbreviation set_difference (infixl "" 65) where "AB A-B" -text‹In ZF set theory the zero of natural numbers is the same as the empty set. - In the next abbreviation we declare that we want $0$ and $\emptyset$ to be synonyms - so that we can use $\emptyset$ instead of $0$ when appropriate. › - -abbreviation empty_set ("") where " 0" end diff --git a/docs/IsarMathLib/document.pdf b/docs/IsarMathLib/document.pdf index cbe9c39..b3285a7 100644 Binary files a/docs/IsarMathLib/document.pdf and b/docs/IsarMathLib/document.pdf differ diff --git a/docs/IsarMathLib/func1.html b/docs/IsarMathLib/func1.html index ce5081a..f8598e3 100644 --- a/docs/IsarMathLib/func1.html +++ b/docs/IsarMathLib/func1.html @@ -245,15 +245,22 @@

Theory func1

lemma func1_1_L6A: assumes A1: "f:XY" shows "f-``(A)X" proof fix x - assume A2: "xf-``(A)" then obtain y where " x,y f" + assume A2: "xf-``(A)" then obtain y where "x,y f" using vimage_iff by auto with A1 show "xX" using func1_1_L5 by fast qed text‹Image of a greater set is greater.› -lemma func1_1_L8: assumes A1: "AB" shows "f``(A) f``(B)" - using assms image_Un by auto +lemma func1_1_L8: assumes A1: "AB" shows "f``(A) f``(B)" + using assms image_Un by auto + +text‹An immediate corollary of vimage_mono› from the Isabelle/ZF distribution - + the inverse image of a greater set is greater. Note we do not require that + $f$ is a function, so this is true for relations as well.› + +lemma vimage_mono1: assumes "AB" shows "f-``(A) f-``(B)" + using assms vimage_mono by simp text‹A set is contained in the the inverse image of its image. There is similar theorem in equalities.thy› @@ -261,57 +268,56 @@

Theory func1

which shows that the image of inverse image of a set is contained in the set.›
-lemma func1_1_L9: assumes A1: "f:XY" and A2: "AX" +lemma func1_1_L9: assumes A1: "f:XY" and A2: "AX" shows "A f-``(f``(A))" proof - - from A1 A2 have "xA. x,f`(x) f" using apply_Pair by auto + from A1 A2 have "xA. x,f`(x) f" using apply_Pair by auto then show ?thesis using image_iff by auto qed text‹The inverse image of the image of the domain is the domain.› -lemma inv_im_dom: assumes A1: "f:XY" shows "f-``(f``(X)) = X" +lemma inv_im_dom: assumes A1: "f:XY" shows "f-``(f``(X)) = X" proof - from A1 show "f-``(f``(X)) X" using func1_1_L3 by simp - from A1 show "X f-``(f``(X))" using func1_1_L9 by simp + from A1 show "f-``(f``(X)) X" using func1_1_L3 by simp + from A1 show "X f-``(f``(X))" using func1_1_L9 by simp qed text‹A technical lemma needed to make the func1_1_L11› proof more clear.› lemma func1_1_L10: - assumes A1: "f X×Y" and A2: "∃!y. (yY x,y f)" + assumes A1: "f X×Y" and A2: "∃!y. (yY x,y f)" shows "∃!y. x,y f" proof - from A2 show "y. x, y f" by auto + from A2 show "y. x, y f" by auto fix y n assume "x,y f" and "x,n f" - with A1 A2 show "y=n" by auto + with A1 A2 show "y=n" by auto qed text‹If $f\subseteq X\times Y$ and for every $x\in X$ there is exactly one $y\in Y$ such that $(x,y)\in f$ then $f$ maps $X$ to $Y$.› -lemma func1_1_L11: +lemma func1_1_L11: assumes "f X×Y" and "xX. ∃!y. yY x,y f" - shows "f: XY" using assms func1_1_L10 Pi_iff_old by simp + shows "f: XY" using assms func1_1_L10 Pi_iff_old by simp -text‹A set defined by a lambda-type expression is a fuction. There is a +text‹A set defined by a lambda-type expression is a function. There is a similar lemma in func.thy, but I had problems with lambda expressions syntax so I could not apply it. This lemma is a workaround for this. Besides, lambda - expressions are not readable. -› + expressions are not readable. › -lemma func1_1_L11A: assumes A1: "xX. b(x) Y" +lemma func1_1_L11A: assumes A1: "xX. b(x) Y" shows "{x,y X×Y. b(x) = y} : XY" proof - let ?f = "{ x,y X×Y. b(x) = y}" have "?f X×Y" by auto moreover have "xX. ∃!y. yY x,y ?f" proof - fix x assume A2: "xX" + fix x assume A2: "xX" show "∃!y. yY x, y {x,y X×Y . b(x) = y}" proof - from A2 A1 show + from A2 A1 show "y. yY x, y {x,y X×Y . b(x) = y}" by simp next @@ -327,14 +333,14 @@

Theory func1

text‹The next lemma will replace func1_1_L11A› one day.› -lemma ZF_fun_from_total: assumes A1: "xX. b(x) Y" +lemma ZF_fun_from_total: assumes A1: "xX. b(x) Y" shows "{x,b(x). xX} : XY" proof - let ?f = "{x,b(x). xX}" - { fix x assume A2: "xX" + { fix x assume A2: "xX" have "∃!y. yY x, y ?f" proof - from A1 A2 show "y. yY x, y ?f" + from A1 A2 show "y. yY x, y ?f" by simp next fix y y1 assume "yY x, y ?f" and "y1Y x, y1 ?f" @@ -342,7 +348,7 @@

Theory func1

qed } then have "xX. ∃!y. yY x,y ?f" by simp - moreover from A1 have "?f X×Y" by auto + moreover from A1 have "?f X×Y" by auto ultimately show ?thesis using func1_1_L11 by simp qed @@ -351,54 +357,54 @@

Theory func1

meta-function (deprecated, use
ZF_fun_from_tot_val(1)› instead).›
lemma func1_1_L11B: - assumes A1: "f:XY" "xX" - and A2: "f = {x,y X×Y. b(x) = y}" + assumes A1: "f:XY" "xX" + and A2: "f = {x,y X×Y. b(x) = y}" shows "f`(x) = b(x)" proof - - from A1 have " x,f`(x) f" using apply_iff by simp - with A2 show ?thesis by simp + from A1 have " x,f`(x) f" using apply_iff by simp + with A2 show ?thesis by simp qed text‹The next lemma will replace func1_1_L11B› one day.› -lemma ZF_fun_from_tot_val: +lemma ZF_fun_from_tot_val: assumes "f:XY" "xX" and "f = {x,b(x). xX}" shows "f`(x) = b(x)" and "b(x)Y" proof - - from assms(1,2) have "x,f`(x) f" using apply_iff by simp - with assms(3) show "f`(x) = b(x)" by simp - from assms(1,2) have "f`(x)Y" by (rule apply_funtype) + from assms(1,2) have "x,f`(x) f" using apply_iff by simp + with assms(3) show "f`(x) = b(x)" by simp + from assms(1,2) have "f`(x)Y" by (rule apply_funtype) with f`(x) = b(x) show "b(x)Y" by simp qed text‹Identical meaning as ZF_fun_from_tot_val›, but phrased a bit differently.› -lemma ZF_fun_from_tot_val0: +lemma ZF_fun_from_tot_val0: assumes "f:XY" and "f = {x,b(x). xX}" shows "xX. f`(x) = b(x)" - using assms ZF_fun_from_tot_val by simp + using assms ZF_fun_from_tot_val by simp text‹Another way of expressing that lambda expression is a function.› -lemma lam_is_fun_range: assumes "f={x,g(x). xX}" +lemma lam_is_fun_range: assumes "f={x,g(x). xX}" shows "f:Xrange(f)" proof - have "xX. g(x) range({x,g(x). xX})" unfolding range_def by auto then have "{x,g(x). xX} : Xrange({x,g(x). xX})" by (rule ZF_fun_from_total) - with assms show ?thesis by auto + with assms show ?thesis by auto qed text‹Yet another way of expressing value of a function.› -lemma ZF_fun_from_tot_val1: +lemma ZF_fun_from_tot_val1: assumes "xX" shows "{x,b(x). xX}`(x)=b(x)" proof - let ?f = "{x,b(x). xX}" have "?f:Xrange(?f)" using lam_is_fun_range by simp - with assms show ?thesis using ZF_fun_from_tot_val0 by simp + with assms show ?thesis using ZF_fun_from_tot_val0 by simp qed text‹An hypotheses-free form of ZF_fun_from_tot_val1›: the value of a function @@ -416,66 +422,66 @@

Theory func1

then $f(x)$ is the empty set. This allows us to conclude that if $y\in f(x)$, then $x$ must be en element of the domain of $f$. ›
-lemma arg_in_domain: assumes "f:XY" "yf`(x)" shows "xX" +lemma arg_in_domain: assumes "f:XY" "yf`(x)" shows "xX" proof - { assume "xX" - with assms have False using func1_1_L1 apply_0 by simp + with assms have False using func1_1_L1 apply_0 by simp } thus ?thesis by auto qed text‹We can extend a function by specifying its values on a set disjoint with the domain.› -lemma func1_1_L11C: assumes A1: "f:XY" and A2: "xA. b(x)B" - and A3: "XA = " and Dg: "g = f {x,b(x). xA}" +lemma func1_1_L11C: assumes A1: "f:XY" and A2: "xA. b(x)B" + and A3: "XA = " and Dg: "g = f {x,b(x). xA}" shows "g : XA YB" "xX. g`(x) = f`(x)" "xA. g`(x) = b(x)" proof - let ?h = "{x,b(x). xA}" - from A1 A2 A3 have - I: "f:XY" "?h : AB" "XA = " + from A1 A2 A3 have + I: "f:XY" "?h : AB" "XA = " using ZF_fun_from_total by auto then have "f?h : XA YB" by (rule fun_disjoint_Un) - with Dg show "g : XA YB" by simp - { fix x assume A4: "xA" - with A1 A3 have "(f?h)`(x) = ?h`(x)" + with Dg show "g : XA YB" by simp + { fix x assume A4: "xA" + with A1 A3 have "(f?h)`(x) = ?h`(x)" using func1_1_L1 fun_disjoint_apply2 by blast - moreover from I A4 have "?h`(x) = b(x)" + moreover from I A4 have "?h`(x) = b(x)" using ZF_fun_from_tot_val by simp ultimately have "(f?h)`(x) = b(x)" by simp - } with Dg show "xA. g`(x) = b(x)" by simp + } with Dg show "xA. g`(x) = b(x)" by simp { fix x assume A5: "xX" - with A3 I have "x domain(?h)" + with A3 I have "x domain(?h)" using func1_1_L1 by auto then have "(f?h)`(x) = f`(x)" using fun_disjoint_apply1 by simp - } with Dg show "xX. g`(x) = f`(x)" by simp + } with Dg show "xX. g`(x) = f`(x)" by simp qed text‹We can extend a function by specifying its value at a point that does not belong to the domain.› -lemma func1_1_L11D: assumes A1: "f:XY" and A2: "aX" - and Dg: "g = f {a,b}" +lemma func1_1_L11D: assumes A1: "f:XY" and A2: "aX" + and Dg: "g = f {a,b}" shows "g : X{a} Y{b}" "xX. g`(x) = f`(x)" "g`(a) = b" proof - let ?h = "{a,b}" - from A1 A2 Dg have I: + from A1 A2 Dg have I: "f:XY" "x{a}. b{b}" "X{a} = " "g = f {x,b. x{a}}" by auto then show "g : X{a} Y{b}" by (rule func1_1_L11C) - from I show "xX. g`(x) = f`(x)" + from I show "xX. g`(x) = f`(x)" by (rule func1_1_L11C) - from I have "x{a}. g`(x) = b" + from I have "x{a}. g`(x) = b" by (rule func1_1_L11C) then show "g`(a) = b" by auto qed @@ -484,10 +490,10 @@

Theory func1

on a set disjoint with the domain and on a point that does not belong to any of those sets.›
-lemma func1_1_L11E: - assumes A1: "f:XY" and - A2: "xA. b(x)B" and - A3: "XA = " and A4: "a XA" +lemma func1_1_L11E: + assumes A1: "f:XY" and + A2: "xA. b(x)B" and + A3: "XA = " and A4: "a XA" and Dg: "g = f {x,b(x). xA} {a,c}" shows "g : XA{a} YB{c}" @@ -496,40 +502,40 @@

Theory func1

"g`(a) = c" proof - let ?h = "f {x,b(x). xA}" - from assms show "g : XA{a} YB{c}" + from assms show "g : XA{a} YB{c}" using func1_1_L11C func1_1_L11D by simp - from A1 A2 A3 have I: + from A1 A2 A3 have I: "f:XY" "xA. b(x)B" "XA = " "?h = f {x,b(x). xA}" by auto - from assms have - II: "?h : XA YB" "a XA" "g = ?h {a,c}" + from assms have + II: "?h : XA YB" "a XA" "g = ?h {a,c}" using func1_1_L11C by auto - then have III: "xXA. g`(x) = ?h`(x)" by (rule func1_1_L11D) - moreover from I have "xX. ?h`(x) = f`(x)" + then have III: "xXA. g`(x) = ?h`(x)" by (rule func1_1_L11D) + moreover from I have "xX. ?h`(x) = f`(x)" by (rule func1_1_L11C) ultimately show "xX. g`(x) = f`(x)" by simp - from I have "xA. ?h`(x) = b(x)" by (rule func1_1_L11C) - with III show "xA. g`(x) = b(x)" by simp - from II show "g`(a) = c" by (rule func1_1_L11D) + from I have "xA. ?h`(x) = b(x)" by (rule func1_1_L11C) + with III show "xA. g`(x) = b(x)" by simp + from II show "g`(a) = c" by (rule func1_1_L11D) qed text‹A way of defining a function on a union of two possibly overlapping sets. We decompose the union into two differences and the intersection and define a function separately on each part.› -lemma fun_union_overlap: assumes "xAB. h(x) Y" "xAB. f(x) Y" "xBA. g(x) Y" +lemma fun_union_overlap: assumes "xAB. h(x) Y" "xAB. f(x) Y" "xBA. g(x) Y" shows "{x,if xAB then f(x) else if xBA then g(x) else h(x). x AB}: AB Y" proof - let ?F = "{x,if xAB then f(x) else if xBA then g(x) else h(x). x AB}" - from assms have "xAB. (if xAB then f(x) else if xBA then g(x) else h(x)) Y" + from assms have "xAB. (if xAB then f(x) else if xBA then g(x) else h(x)) Y" by auto then show ?thesis by (rule ZF_fun_from_total) qed text‹Inverse image of intersection is the intersection of inverse images.› -lemma invim_inter_inter_invim: assumes "f:XY" +lemma invim_inter_inter_invim: assumes "f:XY" shows "f-``(AB) = f-``(A) f-``(B)" - using assms fun_is_fun function_vimage_Int by simp + using assms fun_is_fun function_vimage_Int by simp text‹The inverse image of an intersection of a nonempty collection of sets is the intersection of the @@ -537,53 +543,53 @@

Theory func1

which is proven for the case of two sets.›
lemma func1_1_L12: - assumes A1: "B Pow(Y)" and A2: "B" and A3: "f:XY" + assumes A1: "B Pow(Y)" and A2: "B" and A3: "f:XY" shows "f-``(B) = (UB. f-``(U))" proof - from A2 show "f-``(B) (UB. f-``(U))" by blast + from A2 show "f-``(B) (UB. f-``(U))" by blast show "(UB. f-``(U)) f-``(B)" proof - fix x assume A4: "x (UB. f-``(U))" - from A3 have "UB. f-``(U) X" using func1_1_L6A by simp - with A4 have "UB. xX" by auto - with A2 have "xX" by auto - with A3 have "∃!y. x,y f" using Pi_iff_old by simp - with A2 A4 show "x f-``(B)" using vimage_iff by blast + fix x assume A4: "x (UB. f-``(U))" + from A3 have "UB. f-``(U) X" using func1_1_L6A by simp + with A4 have "UB. xX" by auto + with A2 have "xX" by auto + with A3 have "∃!y. x,y f" using Pi_iff_old by simp + with A2 A4 show "x f-``(B)" using vimage_iff by blast qed qed text‹The inverse image of a set does not change when we intersect the set with the image of the domain.› -lemma inv_im_inter_im: assumes "f:XY" +lemma inv_im_inter_im: assumes "f:XY" shows "f-``(A f``(X)) = f-``(A)" - using assms invim_inter_inter_invim inv_im_dom func1_1_L6A + using assms invim_inter_inter_invim inv_im_dom func1_1_L6A by blast text‹If the inverse image of a set is not empty, then the set is not empty. Proof by contradiction.› -lemma func1_1_L13: assumes A1:"f-``(A) " shows "A" - using assms by auto +lemma func1_1_L13: assumes A1:"f-``(A) " shows "A" + using assms by auto text‹If the image of a set is not empty, then the set is not empty. Proof by contradiction.› -lemma func1_1_L13A: assumes A1: "f``(A)" shows "A" - using assms by auto +lemma func1_1_L13A: assumes A1: "f``(A)" shows "A" + using assms by auto text‹What is the inverse image of a singleton?› -lemma func1_1_L14: assumes "f:XY" +lemma func1_1_L14: assumes "f:XY" shows "f-``({y}) = {xX. f`(x) = y}" - using assms func1_1_L6A vimage_singleton_iff apply_iff by auto + using assms func1_1_L6A vimage_singleton_iff apply_iff by auto text‹A lemma that can be used instead fun_extension_iff› to show that two functions are equal› -lemma func_eq: +lemma func_eq: assumes "f: XY" "g: XZ"and "xX. f`(x) = g`(x)" - shows "f = g" using assms fun_extension_iff by simp + shows "f = g" using assms fun_extension_iff by simp text‹An alternative syntax for defining a function: instead of writing $\{\langle x,p(x)\rangle. x\in X\}$ we can write $\lambda x\in X. p(x)$. › @@ -601,21 +607,21 @@

Theory func1

text‹If a function is equal to an expression $b(x)$ on $X$, then it has to be of the form $\{ \langle x, b(x)\rangle | x\in X\}$. › -lemma func_eq_set_of_pairs: assumes "f:XY" "xX. f`(x) = b(x)" +lemma func_eq_set_of_pairs: assumes "f:XY" "xX. f`(x) = b(x)" shows "f = {x, b(x). x X}" proof - - from assms(1) have "f = {x, f`(x). x X}" using fun_is_set_of_pairs + from assms(1) have "f = {x, f`(x). x X}" using fun_is_set_of_pairs by simp - with assms(2) show ?thesis by simp + with assms(2) show ?thesis by simp qed text‹Function defined on a singleton is a single pair.› -lemma func_singleton_pair: assumes A1: "f : {a}X" +lemma func_singleton_pair: assumes A1: "f : {a}X" shows "f = {a, f`(a)}" proof - let ?g = "{a, f`(a)}" - note A1 + note A1 moreover have "?g : {a} {f`(a)}" using singleton_fun by simp moreover have "x {a}. f`(x) = ?g`(x)" using singleton_apply by simp @@ -625,11 +631,11 @@

Theory func1

text‹A single pair is a function on a singleton. This is similar to singleton_fun› from standard Isabelle/ZF.› -lemma pair_func_singleton: assumes A1: "y Y" +lemma pair_func_singleton: assumes A1: "y Y" shows "{x,y} : {x} Y" proof - have "{x,y} : {x} {y}" using singleton_fun by simp - moreover from A1 have "{y} Y" by simp + moreover from A1 have "{y} Y" by simp ultimately show "{x,y} : {x} Y" by (rule func1_1_L1B) qed @@ -641,67 +647,74 @@

Theory func1

text‹A more familiar definition of inverse image.› -lemma func1_1_L15: assumes A1: "f:XY" +lemma func1_1_L15: assumes A1: "f:XY" shows "f-``(A) = {xX. f`(x) A}" proof - have "f-``(A) = (yA . f-``{y})" by (rule vimage_eq_UN) - with A1 show ?thesis using func1_1_L14 by auto + with A1 show ?thesis using func1_1_L14 by auto qed +text‹For symmetric functions inverse images are symmetric.› + +lemma symm_vimage_symm: + assumes "f:X×XY" and "xX. yX. f`x,y = f`y,x" + shows "f-``(A) = converse(f-``(A))" + using assms func1_1_L15 by auto + text‹A more familiar definition of image.› -lemma func_imagedef: assumes A1: "f:XY" and A2: "AX" +lemma func_imagedef: assumes A1: "f:XY" and A2: "AX" shows "f``(A) = {f`(x). x A}" proof - from A1 show "f``(A) {f`(x). x A}" + from A1 show "f``(A) {f`(x). x A}" using image_iff apply_iff by auto show "{f`(x). x A} f``(A)" proof fix y assume "y {f`(x). x A}" then obtain x where "xA" and "y = f`(x)" by auto - with A1 A2 have "x,y f" using apply_iff by force - with A1 A2 xA show "y f``(A)" using image_iff by auto + with A1 A2 have "x,y f" using apply_iff by force + with A1 A2 xA show "y f``(A)" using image_iff by auto qed qed text‹If all elements of a nonempty set map to the same element of the codomain, then the image of this set is a singleton.› -lemma image_constant_singleton: +lemma image_constant_singleton: assumes "f:XY" "AX" "A" "xA. f`(x) = c" shows "f``(A) = {c}" - using assms func_imagedef by auto + using assms func_imagedef by auto text‹A technical lemma about graphs of functions: if we have two disjoint sets $A$ and $B$ then the cartesian product of the inverse image of $A$ and $B$ is disjoint with (the graph of) $f$.› -lemma vimage_prod_dis_graph: assumes "f:XY" "AB = " +lemma vimage_prod_dis_graph: assumes "f:XY" "AB = " shows "f-``(A)×B f = " proof - { assume "f-``(A)×B f " then obtain p where "p f-``(A)×B" and "pf" by blast - from assms(1) pf have "p {x, f`(x). x X}" + from assms(1) pf have "p {x, f`(x). x X}" using fun_is_set_of_pairs by simp then obtain x where "p = x, f`(x)" by blast - with assms p f-``(A)×B have False using func1_1_L15 by auto + with assms p f-``(A)×B have False using func1_1_L15 by auto } thus ?thesis by auto qed text‹ For two functions with the same domain $X$ and the codomain $Y,Z$ resp., we can define a third one that maps $X$ to the cartesian product of $Y$ and $Z$. › -lemma prod_fun_val: +lemma prod_fun_val: assumes "{x,p(x). xX}: XY" "{x,q(x). xX}: XZ" defines "h {x,p(x),q(x). xX}" shows "h:XY×Z" and "xX. h`(x) = p(x),q(x)" proof - - from assms(1,2) have "xX. p(x),q(x) Y×Z" + from assms(1,2) have "xX. p(x),q(x) Y×Z" using ZF_fun_from_tot_val(2) by auto - with assms(3) show "h:XY×Z" using ZF_fun_from_total by simp - with assms(3) show "xX. h`(x) = p(x),q(x)" using ZF_fun_from_tot_val0 + with assms(3) show "h:XY×Z" using ZF_fun_from_total by simp + with assms(3) show "xX. h`(x) = p(x),q(x)" using ZF_fun_from_tot_val0 by simp qed @@ -712,7 +725,7 @@

Theory func1

$h^{-1}(\{ \langle y,y\rangle : y\in X\}$. It is a bit surprising that we get the last identity without the assumption that $Y=Z$. ›
-lemma vimage_prod: +lemma vimage_prod: assumes "f:XY" "g:XZ" defines "h {x,f`(x),g`(x). xX}" shows @@ -721,121 +734,121 @@

Theory func1

"h-``(U×V) = f-``(U) g-``(V)" "{xX. f`(x) = g`(x)} = h-``({y,y. yY})" proof - - from assms show "h:XY×Z" using apply_funtype ZF_fun_from_total + from assms show "h:XY×Z" using apply_funtype ZF_fun_from_total by simp - with assms(3) show I: "xX. h`(x) = f`(x),g`(x)" + with assms(3) show I: "xX. h`(x) = f`(x),g`(x)" using ZF_fun_from_tot_val by simp - with assms(1,2) h:XY×Z show "h-``(U×V) = f-``(U) g-``(V)" + with assms(1,2) h:XY×Z show "h-``(U×V) = f-``(U) g-``(V)" using func1_1_L15 by auto - from assms(1) I h:XY×Z show "{xX. f`(x) = g`(x)} = h-``({y,y. yY})" + from assms(1) I h:XY×Z show "{xX. f`(x) = g`(x)} = h-``({y,y. yY})" using apply_funtype func1_1_L15 by auto qed text‹The image of a set contained in domain under identity is the same set.› -lemma image_id_same: assumes "AX" shows "id(X)``(A) = A" - using assms id_type id_conv by auto +lemma image_id_same: assumes "AX" shows "id(X)``(A) = A" + using assms id_type id_conv by auto text‹The inverse image of a set contained in domain under identity is the same set.› -lemma vimage_id_same: assumes "AX" shows "id(X)-``(A) = A" - using assms id_type id_conv by auto +lemma vimage_id_same: assumes "AX" shows "id(X)-``(A) = A" + using assms id_type id_conv by auto text‹What is the image of a singleton?› -lemma singleton_image: +lemma singleton_image: assumes "fXY" and "xX" shows "f``{x} = {f`(x)}" - using assms func_imagedef by auto + using assms func_imagedef by auto text‹If an element of the domain of a function belongs to a set, then its value belongs to the image of that set.› -lemma func1_1_L15D: assumes "f:XY" "xA" "AX" +lemma func1_1_L15D: assumes "f:XY" "xA" "AX" shows "f`(x) f``(A)" - using assms func_imagedef by auto + using assms func_imagedef by auto text‹Range is the image of the domain. Isabelle/ZF defines range(f)› as domain(converse(f))›, and that's why we have something to prove here.› lemma range_image_domain: - assumes A1: "f:XY" shows "f``(X) = range(f)" + assumes A1: "f:XY" shows "f``(X) = range(f)" proof show "f``(X) range(f)" using image_def by auto { fix y assume "y range(f)" then obtain x where "y,x converse(f)" by auto - with A1 have "xX" using func1_1_L5 by blast - with A1 have "f`(x) f``(X)" using func_imagedef + with A1 have "xX" using func1_1_L5 by blast + with A1 have "f`(x) f``(X)" using func_imagedef by auto - with A1 y,x converse(f) have "y f``(X)" + with A1 y,x converse(f) have "y f``(X)" using apply_equality by auto } then show "range(f) f``(X)" by auto qed text‹The difference of images is contained in the image of difference.› -lemma diff_image_diff: assumes A1: "f: XY" and A2: "AX" +lemma diff_image_diff: assumes A1: "f: XY" and A2: "AX" shows "f``(X)f``(A) f``(XA)" proof fix y assume "y f``(X) f``(A)" - hence "y f``(X)" and I: "y f``(A)" by auto - with A1 obtain x where "xX" and II: "y = f`(x)" + hence "y f``(X)" and I: "y f``(A)" by auto + with A1 obtain x where "xX" and II: "y = f`(x)" using func_imagedef by auto - with A1 A2 I have "xA" + with A1 A2 I have "xA" using func1_1_L15D by auto with xX have "x XA" "XA X" by auto - with A1 II show "y f``(XA)" + with A1 II show "y f``(XA)" using func1_1_L15D by simp qed text‹The image of an intersection is contained in the intersection of the images.› -lemma image_of_Inter: assumes A1: "f:XY" and - A2: "I" and A3: "iI. P(i) X" +lemma image_of_Inter: assumes A1: "f:XY" and + A2: "I" and A3: "iI. P(i) X" shows "f``(iI. P(i)) ( iI. f``(P(i)) )" proof - fix y assume A4: "y f``(iI. P(i))" - from A1 A2 A3 have "f``(iI. P(i)) = {f`(x). x ( iI. P(i) )}" + fix y assume A4: "y f``(iI. P(i))" + from A1 A2 A3 have "f``(iI. P(i)) = {f`(x). x ( iI. P(i) )}" using ZF1_1_L7 func_imagedef by simp - with A4 obtain x where "x ( iI. P(i) )" and "y = f`(x)" + with A4 obtain x where "x ( iI. P(i) )" and "y = f`(x)" by auto - with A1 A2 A3 show "y ( iI. f``(P(i)) )" using func_imagedef + with A1 A2 A3 show "y ( iI. f``(P(i)) )" using func_imagedef by auto qed text‹The image of union is the union of images.› -lemma image_of_Union: assumes A1: "f:XY" and A2: "AM. AX" +lemma image_of_Union: assumes A1: "f:XY" and A2: "AM. AX" shows "f``(M) = {f``(A). AM}" proof - from A2 have "M X" by auto + from A2 have "M X" by auto { fix y assume "y f``(M)" - with A1 M X obtain x where "xM" and I: "y = f`(x)" + with A1 M X obtain x where "xM" and I: "y = f`(x)" using func_imagedef by auto then obtain A where "AM" and "xA" by auto - with assms I have "y {f``(A). AM}" using func_imagedef by auto + with assms I have "y {f``(A). AM}" using func_imagedef by auto } thus "f``(M) {f``(A). AM}" by auto { fix y assume "y {f``(A). AM}" then obtain A where "AM" and "y f``(A)" by auto - with assms M X have "y f``(M)" using func_imagedef by auto + with assms M X have "y f``(M)" using func_imagedef by auto } thus "{f``(A). AM} f``(M)" by auto qed text‹If the domain of a function is nonempty, then the codomain is as well.› -lemma codomain_nonempty: assumes "f:XY" "X" shows "Y" - using assms apply_funtype by blast +lemma codomain_nonempty: assumes "f:XY" "X" shows "Y" + using assms apply_funtype by blast text‹The image of a nonempty subset of domain is nonempty.› lemma func1_1_L15A: - assumes A1: "f: XY" and A2: "AX" and A3: "A" + assumes A1: "f: XY" and A2: "AX" and A3: "A" shows "f``(A) " proof - - from A3 obtain x where "xA" by auto - with A1 A2 have "f`(x) f``(A)" + from A3 obtain x where "xA" by auto + with A1 A2 have "f`(x) f``(A)" using func_imagedef by auto then show "f``(A) " by auto qed @@ -843,60 +856,60 @@

Theory func1

text‹The next lemma allows to prove statements about the values in the domain of a function given a statement about values in the range.› -lemma func1_1_L15B: +lemma func1_1_L15B: assumes "f:XY" and "AX" and "yf``(A). P(y)" shows "xA. P(f`(x))" - using assms func_imagedef by simp + using assms func_imagedef by simp text‹An image of an image is the image of a composition.› -lemma func1_1_L15C: assumes A1: "f:XY" and A2: "g:YZ" - and A3: "AX" +lemma func1_1_L15C: assumes A1: "f:XY" and A2: "g:YZ" + and A3: "AX" shows "g``(f``(A)) = {g`(f`(x)). xA}" "g``(f``(A)) = (g O f)``(A)" proof - - from A1 A3 have "{f`(x). xA} Y" + from A1 A3 have "{f`(x). xA} Y" using apply_funtype by auto - with A2 have "g``{f`(x). xA} = {g`(f`(x)). xA}" + with A2 have "g``{f`(x). xA} = {g`(f`(x)). xA}" using func_imagedef by auto - with A1 A3 show I: "g``(f``(A)) = {g`(f`(x)). xA}" + with A1 A3 show I: "g``(f``(A)) = {g`(f`(x)). xA}" using func_imagedef by simp - from A1 A3 have "xA. (g O f)`(x) = g`(f`(x))" + from A1 A3 have "xA. (g O f)`(x) = g`(f`(x))" using comp_fun_apply by auto - with I have "g``(f``(A)) = {(g O f)`(x). xA}" + with I have "g``(f``(A)) = {(g O f)`(x). xA}" by simp - moreover from A1 A2 A3 have "(g O f)``(A) = {(g O f)`(x). xA}" + moreover from A1 A2 A3 have "(g O f)``(A) = {(g O f)`(x). xA}" using comp_fun func_imagedef by blast ultimately show "g``(f``(A)) = (g O f)``(A)" by simp qed -text‹What is the image of a set defined by a meta-fuction?› +text‹What is the image of a set defined by a meta-function?› lemma func1_1_L17: - assumes A1: "f XY" and A2: "xA. b(x) X" + assumes A1: "f XY" and A2: "xA. b(x) X" shows "f``({b(x). xA}) = {f`(b(x)). xA}" proof - - from A2 have "{b(x). xA} X" by auto - with A1 show ?thesis using func_imagedef by auto + from A2 have "{b(x). xA} X" by auto + with A1 show ?thesis using func_imagedef by auto qed text‹What are the values of composition of three functions?› -lemma func1_1_L18: assumes A1: "f:AB" "g:BC" "h:CD" - and A2: "xA" +lemma func1_1_L18: assumes A1: "f:AB" "g:BC" "h:CD" + and A2: "xA" shows "(h O g O f)`(x) D" "(h O g O f)`(x) = h`(g`(f`(x)))" proof - - from A1 have "(h O g O f) : AD" + from A1 have "(h O g O f) : AD" using comp_fun by blast - with A2 show "(h O g O f)`(x) D" using apply_funtype + with A2 show "(h O g O f)`(x) D" using apply_funtype by simp - from A1 A2 have "(h O g O f)`(x) = h`( (g O f)`(x))" + from A1 A2 have "(h O g O f)`(x) = h`( (g O f)`(x))" using comp_fun comp_fun_apply by blast - with A1 A2 show "(h O g O f)`(x) = h`(g`(f`(x)))" + with A1 A2 show "(h O g O f)`(x) = h`(g`(f`(x)))" using comp_fun_apply by simp qed @@ -904,24 +917,24 @@

Theory func1

generalization of standard Isabelle's
comp_fun›. › lemma comp_fun_subset: - assumes A1: "g:AB" and A2: "f:CD" and A3: "B C" + assumes A1: "g:AB" and A2: "f:CD" and A3: "B C" shows "f O g : A D" proof - - from A1 A3 have "g:AC" by (rule func1_1_L1B) - with A2 show "f O g : A D" using comp_fun by simp + from A1 A3 have "g:AC" by (rule func1_1_L1B) + with A2 show "f O g : A D" using comp_fun by simp qed text‹This lemma supersedes the lemma comp_eq_id_iff› in Isabelle/ZF. Contributed by Victor Porton.› -lemma comp_eq_id_iff1: assumes A1: "g: BA" and A2: "f: AC" +lemma comp_eq_id_iff1: assumes A1: "g: BA" and A2: "f: AC" shows "(yB. f`(g`(y)) = y) f O g = id(B)" proof - - from assms have "f O g: BC" and "id(B): BB" + from assms have "f O g: BC" and "id(B): BB" using comp_fun id_type by auto then have "(yB. (f O g)`y = id(B)`(y)) f O g = id(B)" by (rule fun_extension_iff) - moreover from A1 have + moreover from A1 have "yB. (f O g)`y = f`(g`y)" and "yB. id(B)`(y) = y" by auto ultimately show "(yB. f`(g`y) = y) f O g = id(B)" by simp @@ -930,14 +943,14 @@

Theory func1

text‹A lemma about a value of a function that is a union of some collection of functions.› -lemma fun_Union_apply: assumes A1: "F : XY" and - A2: "fF" and A3: "f:AB" and A4: "xA" +lemma fun_Union_apply: assumes A1: "F : XY" and + A2: "fF" and A3: "f:AB" and A4: "xA" shows "(F)`(x) = f`(x)" proof - - from A3 A4 have "x, f`(x) f" using apply_Pair + from A3 A4 have "x, f`(x) f" using apply_Pair by simp - with A2 have "x, f`(x) F" by auto - with A1 show "(F)`(x) = f`(x)" using apply_equality + with A2 have "x, f`(x) F" by auto + with A1 show "(F)`(x) = f`(x)" using apply_equality by simp qed @@ -951,13 +964,13 @@

Theory func1

text‹What is the inverse image of a set under a restricted fuction?› -lemma func1_2_L1: assumes A1: "f:XY" and A2: "BX" +lemma func1_2_L1: assumes A1: "f:XY" and A2: "BX" shows "restrict(f,B)-``(A) = f-``(A) B" proof - let ?g = "restrict(f,B)" - from A1 A2 have "?g:BY" + from A1 A2 have "?g:BY" using restrict_type2 by simp - with A2 A1 show "?g-``(A) = f-``(A) B" + with A2 A1 show "?g-``(A) = f-``(A) B" using func1_1_L15 restrict_if by auto qed @@ -966,85 +979,85 @@

Theory func1

criterion and applications.›
lemma func1_2_L2: - assumes A1: "f:XY" and A2: "g AZ" - and A3: "AX" and A4: "f A×Z = g" + assumes A1: "f:XY" and A2: "g AZ" + and A3: "AX" and A4: "f A×Z = g" shows "xA. g`(x) = f`(x)" proof fix x assume "xA" - with A2 have "x,g`(x) g" using apply_Pair by simp - with A4 A1 show "g`(x) = f`(x)" using apply_iff by auto + with A2 have "x,g`(x) g" using apply_Pair by simp + with A4 A1 show "g`(x) = f`(x)" using apply_iff by auto qed text‹Here is the actual criterion.› lemma func1_2_L3: - assumes A1: "f:XY" and A2: "g:AZ" - and A3: "AX" and A4: "f A×Z = g" + assumes A1: "f:XY" and A2: "g:AZ" + and A3: "AX" and A4: "f A×Z = g" shows "g = restrict(f,A)" proof - from A4 show "g restrict(f, A)" using restrict_iff by auto + from A4 show "g restrict(f, A)" using restrict_iff by auto show "restrict(f, A) g" proof fix z assume A5:"z restrict(f,A)" - then obtain x y where D1:"zf xA z = x,y" + then obtain x y where D1:"zf xA z = x,y" using restrict_iff by auto - with A1 have "y = f`(x)" using apply_iff by auto - with A1 A2 A3 A4 D1 have "y = g`(x)" using func1_2_L2 by simp - with A2 D1 show "zg" using apply_Pair by simp + with A1 have "y = f`(x)" using apply_iff by auto + with A1 A2 A3 A4 D1 have "y = g`(x)" using func1_2_L2 by simp + with A2 D1 show "zg" using apply_Pair by simp qed qed text‹Which function space a restricted function belongs to?› lemma func1_2_L4: - assumes A1: "f:XY" and A2: "AX" and A3: "xA. f`(x) Z" + assumes A1: "f:XY" and A2: "AX" and A3: "xA. f`(x) Z" shows "restrict(f,A) : AZ" proof - let ?g = "restrict(f,A)" - from A1 A2 have "?g : AY" + from A1 A2 have "?g : AY" using restrict_type2 by simp moreover { fix x assume "xA" - with A1 A3 have "?g`(x) Z" using restrict by simp} + with A1 A3 have "?g`(x) Z" using restrict by simp} ultimately show ?thesis by (rule Pi_type) qed text‹A simpler case of func1_2_L4›, where the range of the original and restricted function are the same.› -corollary restrict_fun: assumes A1: "f:XY" and A2: "AX" +corollary restrict_fun: assumes A1: "f:XY" and A2: "AX" shows "restrict(f,A) : A Y" proof - - from assms have "xA. f`(x) Y" using apply_funtype + from assms have "xA. f`(x) Y" using apply_funtype by auto - with assms show ?thesis using func1_2_L4 by simp + with assms show ?thesis using func1_2_L4 by simp qed text‹A function restricted to its domain is itself.› -lemma restrict_domain: assumes "f:XY" +lemma restrict_domain: assumes "f:XY" shows "restrict(f,X) = f" proof - have "xX. restrict(f,X)`(x) = f`(x)" using restrict by simp - with assms show ?thesis using func_eq restrict_fun by blast + with assms show ?thesis using func_eq restrict_fun by blast qed text‹Suppose a function $f:X\rightarrow Y$ is defined by an expression $q$, i.e. $f = \{\langle x,y\rangle : x\in X\}$. Then a function that is defined by the same expression, but on a smaller set is the same as the restriction of $f$ to that smaller set.› -lemma restrict_def_alt: assumes "AX" +lemma restrict_def_alt: assumes "AX" shows "restrict({x,q(x). xX},A) = {x,q(x). xA}" proof - let ?Y = "{q(x). xX}" let ?f = "{x,q(x). xX}" have "xX. q(x)?Y" by blast - with assms have "?f:X?Y" using ZF_fun_from_total by simp - with assms have "restrict(?f,A):A?Y" using restrict_fun by simp + with assms have "?f:X?Y" using ZF_fun_from_total by simp + with assms have "restrict(?f,A):A?Y" using restrict_fun by simp moreover - from assms have "xA. q(x)?Y" by blast + from assms have "xA. q(x)?Y" by blast then have "{x,q(x). xA}:A?Y" using ZF_fun_from_total by simp - moreover from assms have + moreover from assms have "xA. restrict(?f,A)`(x) = {x,q(x). xA}`(x)" using restrict ZF_fun_from_tot_val1 by auto ultimately show ?thesis by (rule func_eq) @@ -1053,15 +1066,15 @@

Theory func1

text‹A composition of two functions is the same as composition with a restriction.› -lemma comp_restrict: - assumes A1: "f : AB" and A2: "g : X C" and A3: "BX" +lemma comp_restrict: + assumes A1: "f : AB" and A2: "g : X C" and A3: "BX" shows "g O f = restrict(g,B) O f" proof - - from assms have "g O f : A C" using comp_fun_subset + from assms have "g O f : A C" using comp_fun_subset by simp - moreover from assms have "restrict(g,B) O f : A C" + moreover from assms have "restrict(g,B) O f : A C" using restrict_fun comp_fun by simp - moreover from A1 have + moreover from A1 have "xA. (g O f)`(x) = (restrict(g,B) O f)`(x)" using comp_fun_apply apply_funtype restrict by simp @@ -1088,22 +1101,22 @@

Theory func1

text‹Constant function is a function (i.e. belongs to a function space).› lemma func1_3_L1: - assumes A1: "cY" shows "ConstantFunction(X,c) : XY" + assumes A1: "cY" shows "ConstantFunction(X,c) : XY" proof - - from A1 have "X×{c} = { x,y X×Y. c = y}" + from A1 have "X×{c} = { x,y X×Y. c = y}" by auto - with A1 show ?thesis using func1_1_L11A ConstantFunction_def + with A1 show ?thesis using func1_1_L11A ConstantFunction_def by simp qed text‹Constant function is equal to the constant on its domain.› -lemma func1_3_L2: assumes A1: "xX" +lemma func1_3_L2: assumes A1: "xX" shows "ConstantFunction(X,c)`(x) = c" proof - have "ConstantFunction(X,c) X{c}" using func1_3_L1 by simp - moreover from A1 have "x,c ConstantFunction(X,c)" + moreover from A1 have "x,c ConstantFunction(X,c)" using ConstantFunction_def by simp ultimately show ?thesis using apply_iff by simp qed @@ -1117,25 +1130,25 @@

Theory func1

text‹If $c\in A$ then the inverse image of $A$ by the constant function $x\mapsto c$ is the whole domain. › -lemma const_vimage_domain: assumes "cA" +lemma const_vimage_domain: assumes "cA" shows "ConstantFunction(X,c)-``(A) = X" proof - let ?C = "ConstantFunction(X,c)" have "?C-``(A) = {xX. ?C`(x) A}" using func1_3_L1 func1_1_L15 by blast - with assms show ?thesis using func1_3_L2 by simp + with assms show ?thesis using func1_3_L2 by simp qed text‹If $c$ is not an element of $A$ then the inverse image of $A$ by the constant function $x\mapsto c$ is empty. › -lemma const_vimage_empty: assumes "cA" +lemma const_vimage_empty: assumes "cA" shows "ConstantFunction(X,c)-``(A) = " proof - let ?C = "ConstantFunction(X,c)" have "?C-``(A) = {xX. ?C`(x) A}" using func1_3_L1 func1_1_L15 by blast - with assms show ?thesis using func1_3_L2 by simp + with assms show ?thesis using func1_3_L2 by simp qed subsection‹Injections, surjections, bijections etc.› @@ -1148,28 +1161,28 @@

Theory func1

the difference of images›
lemma inj_image_dif: - assumes A1: "f inj(A,B)" and A2: "C A" + assumes A1: "f inj(A,B)" and A2: "C A" shows "f``(AC) = f``(A)f``(C)" proof show "f``(AC) f``(A)f``(C)" proof - fix y assume A3: "y f``(AC)" - from A1 have "f:AB" using inj_def by simp + fix y assume A3: "y f``(AC)" + from A1 have "f:AB" using inj_def by simp moreover have "AC A" by auto ultimately have "f``(AC) = {f`(x). x AC}" using func_imagedef by simp - with A3 obtain x where I: "f`(x) = y" and "x AC" + with A3 obtain x where I: "f`(x) = y" and "x AC" by auto hence "xA" by auto - with f:AB I have "y f``(A)" + with f:AB I have "y f``(A)" using func_imagedef by auto moreover have "y f``(C)" proof - { assume "y f``(C)" - with A2 f:AB obtain x0 + with A2 f:AB obtain x0 where II: "f`(x0) = y" and "x0 C" using func_imagedef by auto - with A1 A2 I xA have + with A1 A2 I xA have "f inj(A,B)" "f`(x) = f`(x0)" "xA" "x0 A" by auto then have "x = x0" by (rule inj_apply_equality) @@ -1178,41 +1191,41 @@

Theory func1

qed ultimately show "y f``(A)f``(C)" by simp qed - from A1 A2 show "f``(A)f``(C) f``(AC)" + from A1 A2 show "f``(A)f``(C) f``(AC)" using inj_def diff_image_diff by auto qed text‹For injections the image of intersection is the intersection of images.› -lemma inj_image_inter: assumes A1: "f inj(X,Y)" and A2: "AX" "BX" +lemma inj_image_inter: assumes A1: "f inj(X,Y)" and A2: "AX" "BX" shows "f``(AB) = f``(A) f``(B)" proof show "f``(AB) f``(A) f``(B)" using image_Int_subset by simp - { from A1 have "f:XY" using inj_def by simp + { from A1 have "f:XY" using inj_def by simp fix y assume "y f``(A) f``(B)" then have "y f``(A)" and "y f``(B)" by auto - with A2 f:XY obtain xA xB where - "xA A" "xB B" and I: "y = f`(xA)" "y = f`(xB)" + with A2 f:XY obtain xA xB where + "xA A" "xB B" and I: "y = f`(xA)" "y = f`(xB)" using func_imagedef by auto - with A2 have "xA X" "xB X" and " f`(xA) = f`(xB)" by auto - with A1 have "xA = xB" using inj_def by auto + with A2 have "xA X" "xB X" and " f`(xA) = f`(xB)" by auto + with A1 have "xA = xB" using inj_def by auto with xA A xB B have "f`(xA) {f`(x). x AB}" by auto - moreover from A2 f:XY have "f``(AB) = {f`(x). x AB}" + moreover from A2 f:XY have "f``(AB) = {f`(x). x AB}" using func_imagedef by blast ultimately have "f`(xA) f``(AB)" by simp - with I have "y f``(AB)" by simp + with I have "y f``(AB)" by simp } thus "f``(A) f``(B) f``(A B)" by auto qed text‹For surjection from $A$ to $B$ the image of the domain is $B$.› -lemma surj_range_image_domain: assumes A1: "f surj(A,B)" +lemma surj_range_image_domain: assumes A1: "f surj(A,B)" shows "f``(A) = B" proof - - from A1 have "f``(A) = range(f)" + from A1 have "f``(A) = range(f)" using surj_def range_image_domain by auto - with A1 show "f``(A) = B" using surj_range + with A1 show "f``(A) = B" using surj_range by simp qed @@ -1236,7 +1249,7 @@

Theory func1

using inj_image_dif surj_range_image_domain surj_is_fun unfolding bij_def by auto { fix f assume "f?R" - hence "f:XY" and I: "APow(X). f``(XA) = Yf``(A)" + hence "f:XY" and I: "APow(X). f``(XA) = Yf``(A)" by auto { fix x1 x2 assume "x1X" "x2X" "f`(x1) = f`(x2)" with f:XY have @@ -1245,14 +1258,14 @@

Theory func1

{ assume "x1x2" from f:XY have "f``(X{x1}) = {f`(t). tX{x1}}" using func_imagedef by blast - with I x2X x1X x1x2 f``{x1} = f``{x2} + with I x2X x1X x1x2 f``{x1} = f``{x2} have "f`(x2) Yf``{x2}" by auto with f``{x2} = {f`(x2)} have False by auto } hence "x1=x2" by auto } with f:XY have "finj(X,Y)" unfolding inj_def by auto moreover - from I have "f``(X) = Yf``()" by blast + from I have "f``(X) = Yf``()" by blast with f:XY have "fsurj(X,Y)" using surj_def_alt by simp ultimately have "f bij(X,Y)" unfolding bij_def by simp } thus "?R bij(X,Y)" by auto @@ -1260,43 +1273,43 @@

Theory func1

text‹For injections the inverse image of an image is the same set.› -lemma inj_vimage_image: assumes "f inj(X,Y)" and "AX" +lemma inj_vimage_image: assumes "f inj(X,Y)" and "AX" shows "f-``(f``(A)) = A" proof - have "f-``(f``(A)) = (converse(f) O f)``(A)" using vimage_converse image_comp by simp - with assms show ?thesis using left_comp_inverse image_id_same + with assms show ?thesis using left_comp_inverse image_id_same by simp qed text‹For surjections the image of an inverse image is the same set.› -lemma surj_image_vimage: assumes A1: "f surj(X,Y)" and A2: "AY" +lemma surj_image_vimage: assumes A1: "f surj(X,Y)" and A2: "AY" shows "f``(f-``(A)) = A" proof - have "f``(f-``(A)) = (f O converse(f))``(A)" using vimage_converse image_comp by simp - with assms show ?thesis using right_comp_inverse image_id_same + with assms show ?thesis using right_comp_inverse image_id_same by simp qed text‹A lemma about how a surjection maps collections of subsets in domain and range.› -lemma surj_subsets: assumes A1: "f surj(X,Y)" and A2: "B Pow(Y)" +lemma surj_subsets: assumes A1: "f surj(X,Y)" and A2: "B Pow(Y)" shows "{ f``(U). U {f-``(V). VB} } = B" proof { fix W assume "W { f``(U). U {f-``(V). VB} }" - then obtain U where I: "U {f-``(V). VB}" and II: "W = f``(U)" by auto + then obtain U where I: "U {f-``(V). VB}" and II: "W = f``(U)" by auto then obtain V where "VB" and "U = f-``(V)" by auto - with II have "W = f``(f-``(V))" by simp - moreover from assms VB have "f surj(X,Y)" and "VY" by auto + with II have "W = f``(f-``(V))" by simp + moreover from assms VB have "f surj(X,Y)" and "VY" by auto ultimately have "W=V" using surj_image_vimage by simp with VB have "W B" by simp } thus "{ f``(U). U {f-``(V). VB} } B" by auto { fix W assume "WB" let ?U = "f-``(W)" from WB have "?U {f-``(V). VB}" by auto - moreover from A1 A2 WB have "W = f``(?U)" using surj_image_vimage by auto + moreover from A1 A2 WB have "W = f``(?U)" using surj_image_vimage by auto ultimately have "W { f``(U). U {f-``(V). VB} }" by auto } thus "B { f``(U). U {f-``(V). VB} }" by auto qed @@ -1305,21 +1318,21 @@

Theory func1

is a a bijection.›
lemma bij_restrict_rem: - assumes A1: "f bij(A,B)" and A2: "aA" + assumes A1: "f bij(A,B)" and A2: "aA" shows "restrict(f, A{a}) bij(A{a}, B{f`(a)})" proof - let ?C = "A{a}" - from A1 have "f inj(A,B)" "?C A" + from A1 have "f inj(A,B)" "?C A" using bij_def by auto then have "restrict(f,?C) bij(?C, f``(?C))" using restrict_bij by simp moreover have "f``(?C) = B{f`(a)}" proof - - from A2 f inj(A,B) have "f``(?C) = f``(A)f``{a}" + from A2 f inj(A,B) have "f``(?C) = f``(A)f``{a}" using inj_image_dif by simp - moreover from A1 have "f``(A) = B" + moreover from A1 have "f``(A) = B" using bij_def surj_range_image_domain by auto - moreover from A1 A2 have "f``{a} = {f`(a)}" + moreover from A1 A2 have "f``{a} = {f`(a)}" using bij_is_fun singleton_image by blast ultimately show "f``(?C) = B{f`(a)}" by simp qed @@ -1329,9 +1342,9 @@

Theory func1

text‹The domain of a bijection between $X$ and $Y$ is $X$.› lemma domain_of_bij: - assumes A1: "f bij(X,Y)" shows "domain(f) = X" + assumes A1: "f bij(X,Y)" shows "domain(f) = X" proof - - from A1 have "f:XY" using bij_is_fun by simp + from A1 have "f:XY" using bij_is_fun by simp then show "domain(f) = X" using func1_1_L1 by simp qed @@ -1339,17 +1352,17 @@

Theory func1

of a set belongs to that set.›
lemma inj_inv_back_in_set: - assumes A1: "f inj(A,B)" and A2: "CA" and A3: "y f``(C)" + assumes A1: "f inj(A,B)" and A2: "CA" and A3: "y f``(C)" shows "converse(f)`(y) C" "f`(converse(f)`(y)) = y" proof - - from A1 have I: "f:AB" using inj_is_fun by simp - with A2 A3 obtain x where II: "xC" "y = f`(x)" + from A1 have I: "f:AB" using inj_is_fun by simp + with A2 A3 obtain x where II: "xC" "y = f`(x)" using func_imagedef by auto - with A1 A2 show "converse(f)`(y) C" using left_inverse + with A1 A2 show "converse(f)`(y) C" using left_inverse by auto - from A1 A2 I II show "f`(converse(f)`(y)) = y" + from A1 A2 I II show "f`(converse(f)`(y)) = y" using func1_1_L5A right_inverse by auto qed @@ -1361,17 +1374,17 @@

Theory func1

of the inverse function) $f^{-1}(3) = \emptyset = 0 \in \{ 0,1\}$, but $3$ is not in the image $f(\{ 0,1\})$. ›
-lemma bij_val_image_vimage: assumes "f bij(X,Y)" "AX" "yY" +lemma bij_val_image_vimage: assumes "f bij(X,Y)" "AX" "yY" shows "yf``(A) converse(f)`(y) A" proof assume "yf``(A)" - with assms(1,2) show "converse(f)`(y) A" + with assms(1,2) show "converse(f)`(y) A" unfolding bij_def using inj_inv_back_in_set by blast next assume "converse(f)`(y) A" - with assms(1,3) have "y{f`(x). xA}" using right_inverse_bij + with assms(1,3) have "y{f`(x). xA}" using right_inverse_bij by force - with assms(1,2) show "yf``(A)" using bij_is_fun func_imagedef + with assms(1,2) show "yf``(A)" using bij_is_fun func_imagedef by force qed @@ -1380,13 +1393,13 @@

Theory func1

belongs to the set.›
lemma inj_point_of_image: - assumes A1: "f inj(A,B)" and A2: "CA" and - A3: "xA" and A4: "f`(x) f``(C)" + assumes A1: "f inj(A,B)" and A2: "CA" and + A3: "xA" and A4: "f`(x) f``(C)" shows "x C" proof - - from A1 A2 A4 have "converse(f)`(f`(x)) C" + from A1 A2 A4 have "converse(f)`(f`(x)) C" using inj_inv_back_in_set by simp - moreover from A1 A3 have "converse(f)`(f`(x)) = x" + moreover from A1 A3 have "converse(f)`(f`(x)) = x" using left_inverse_eq by simp ultimately show "x C" by simp qed @@ -1394,62 +1407,62 @@

Theory func1

text‹For injections the image of intersection is the intersection of images.› -lemma inj_image_of_Inter: assumes A1: "f inj(A,B)" and - A2: "I" and A3: "iI. P(i) A" +lemma inj_image_of_Inter: assumes A1: "f inj(A,B)" and + A2: "I" and A3: "iI. P(i) A" shows "f``(iI. P(i)) = ( iI. f``(P(i)) )" proof - from A1 A2 A3 show "f``(iI. P(i)) ( iI. f``(P(i)) )" + from A1 A2 A3 show "f``(iI. P(i)) ( iI. f``(P(i)) )" using inj_is_fun image_of_Inter by auto - from A1 A2 A3 have "f:AB" and "( iI. P(i) ) A" + from A1 A2 A3 have "f:AB" and "( iI. P(i) ) A" using inj_is_fun ZF1_1_L7 by auto - then have I: "f``(iI. P(i)) = { f`(x). x ( iI. P(i) ) }" + then have I: "f``(iI. P(i)) = { f`(x). x ( iI. P(i) ) }" using func_imagedef by simp - { fix y assume A4: "y ( iI. f``(P(i)) )" + { fix y assume A4: "y ( iI. f``(P(i)) )" let ?x = "converse(f)`(y)" - from A2 obtain i0 where "i0 I" by auto - with A1 A4 have II: "y range(f)" using inj_is_fun func1_1_L6 + from A2 obtain i0 where "i0 I" by auto + with A1 A4 have II: "y range(f)" using inj_is_fun func1_1_L6 by auto - with A1 have III: "f`(?x) = y" using right_inverse by simp - from A1 II have IV: "?x A" using inj_converse_fun apply_funtype + with A1 have III: "f`(?x) = y" using right_inverse by simp + from A1 II have IV: "?x A" using inj_converse_fun apply_funtype by blast { fix i assume "iI" - with A3 A4 III have "P(i) A" and "f`(?x) f``(P(i))" + with A3 A4 III have "P(i) A" and "f`(?x) f``(P(i))" by auto - with A1 IV have "?x P(i)" using inj_point_of_image + with A1 IV have "?x P(i)" using inj_point_of_image by blast } then have "iI. ?x P(i)" by simp - with A2 I have "f`(?x) f``( iI. P(i) )" + with A2 I have "f`(?x) f``( iI. P(i) )" by auto - with III have "y f``( iI. P(i) )" by simp + with III have "y f``( iI. P(i) )" by simp } then show "( iI. f``(P(i)) ) f``( iI. P(i) )" by auto qed text‹An injection is injective onto its range. Suggested by Victor Porton.› -lemma inj_inj_range: assumes "f inj(A,B)" +lemma inj_inj_range: assumes "f inj(A,B)" shows "f inj(A,range(f))" - using assms inj_def range_of_fun by auto + using assms inj_def range_of_fun by auto text‹An injection is a bijection on its range. Suggested by Victor Porton.› -lemma inj_bij_range: assumes "f inj(A,B)" +lemma inj_bij_range: assumes "f inj(A,B)" shows "f bij(A,range(f))" proof - - from assms have "f surj(A,range(f))" using inj_def fun_is_surj + from assms have "f surj(A,range(f))" using inj_def fun_is_surj by auto - with assms show ?thesis using inj_inj_range bij_def by simp + with assms show ?thesis using inj_inj_range bij_def by simp qed text‹A lemma about extending a surjection by one point.› lemma surj_extend_point: - assumes A1: "f surj(X,Y)" and A2: "aX" and - A3: "g = f {a,b}" + assumes A1: "f surj(X,Y)" and A2: "aX" and + A3: "g = f {a,b}" shows "g surj(X{a},Y{b})" proof - - from A1 A2 A3 have "g : X{a} Y{b}" + from A1 A2 A3 have "g : X{a} Y{b}" using surj_def func1_1_L11D by simp moreover have "y Y{b}. x X{a}. y = g`(x)" proof @@ -1457,14 +1470,14 @@

Theory func1

then have "y Y y = b" by auto moreover { assume "y Y" - with A1 obtain x where "xX" and "y = f`(x)" + with A1 obtain x where "xX" and "y = f`(x)" using surj_def by auto - with A1 A2 A3 have "x X{a}" and "y = g`(x)" + with A1 A2 A3 have "x X{a}" and "y = g`(x)" using surj_def func1_1_L11D by auto then have "x X{a}. y = g`(x)" by auto } moreover { assume "y = b" - with A1 A2 A3 have "y = g`(a)" + with A1 A2 A3 have "y = g`(a)" using surj_def func1_1_L11D by auto then have "x X{a}. y = g`(x)" by auto } ultimately show "x X{a}. y = g`(x)" @@ -1478,11 +1491,11 @@

Theory func1

Essentially the same as standard Isabelle's
inj_extend›. › -lemma inj_extend_point: assumes "f inj(X,Y)" "aX" "bY" +lemma inj_extend_point: assumes "f inj(X,Y)" "aX" "bY" shows "(f {a,b}) inj(X{a},Y{b})" proof - - from assms have "cons(a,b,f) inj(cons(a, X), cons(b, Y))" - using assms inj_extend by simp + from assms have "cons(a,b,f) inj(cons(a, X), cons(b, Y))" + using assms inj_extend by simp moreover have "cons(a,b,f) = f {a,b}" and "cons(a, X) = X{a}" and "cons(b, Y) = Y{b}" by auto @@ -1491,25 +1504,25 @@

Theory func1

text‹A lemma about extending a bijection by one point.› -lemma bij_extend_point: assumes "f bij(X,Y)" "aX" "bY" +lemma bij_extend_point: assumes "f bij(X,Y)" "aX" "bY" shows "(f {a,b}) bij(X{a},Y{b})" - using assms surj_extend_point inj_extend_point bij_def + using assms surj_extend_point inj_extend_point bij_def by simp text‹A quite general form of the $a^{-1}b = 1$ implies $a=b$ law.› lemma comp_inv_id_eq: - assumes A1: "converse(b) O a = id(A)" and - A2: "a A×B" "b surj(A,B)" + assumes A1: "converse(b) O a = id(A)" and + A2: "a A×B" "b surj(A,B)" shows "a = b" proof - - from A1 have "(b O converse(b)) O a = b O id(A)" + from A1 have "(b O converse(b)) O a = b O id(A)" using comp_assoc by simp - with A2 have "id(B) O a = b O id(A)" + with A2 have "id(B) O a = b O id(A)" using right_comp_inverse by simp moreover - from A2 have "a A×B" and "b A×B" + from A2 have "a A×B" and "b A×B" using surj_def fun_subset_prod by auto then have "id(B) O a = a" and "b O id(A) = b" @@ -1521,46 +1534,46 @@

Theory func1

the $a^{-1}b = 1$ implies $a=b$ law for bijections.›
lemma comp_inv_id_eq_bij: - assumes A1: "a bij(A,B)" "b bij(A,B)" and - A2: "converse(b) O a = id(A)" + assumes A1: "a bij(A,B)" "b bij(A,B)" and + A2: "converse(b) O a = id(A)" shows "a = b" proof - - from A1 have "a A×B" and "b surj(A,B)" + from A1 have "a A×B" and "b surj(A,B)" using bij_def surj_def fun_subset_prod by auto - with A2 show "a = b" by (rule comp_inv_id_eq) + with A2 show "a = b" by (rule comp_inv_id_eq) qed text‹Converse of a converse of a bijection is the same bijection. This is a special case of converse_converse› from standard Isabelle's equalities› theory where it is proved for relations.› -lemma bij_converse_converse: assumes "a bij(A,B)" +lemma bij_converse_converse: assumes "a bij(A,B)" shows "converse(converse(a)) = a" proof - - from assms have "a A×B" using bij_def surj_def fun_subset_prod by simp + from assms have "a A×B" using bij_def surj_def fun_subset_prod by simp then show ?thesis using converse_converse by simp qed text‹If a composition of bijections is identity, then one is the inverse of the other.› -lemma comp_id_conv: assumes A1: "a bij(A,B)" "b bij(B,A)" and +lemma comp_id_conv: assumes A1: "a bij(A,B)" "b bij(B,A)" and A2: "b O a = id(A)" shows "a = converse(b)" and "b = converse(a)" proof - - from A1 have "a bij(A,B)" and "converse(b) bij(A,B)" using bij_converse_bij + from A1 have "a bij(A,B)" and "converse(b) bij(A,B)" using bij_converse_bij by auto - moreover from assms have "converse(converse(b)) O a = id(A)" + moreover from assms have "converse(converse(b)) O a = id(A)" using bij_converse_converse by simp ultimately show "a = converse(b)" by (rule comp_inv_id_eq_bij) - with assms show "b = converse(a)" using bij_converse_converse by simp + with assms show "b = converse(a)" using bij_converse_converse by simp qed text‹A version of comp_id_conv› with weaker assumptions.› -lemma comp_conv_id: assumes A1: "a bij(A,B)" and A2: "b:BA" and - A3: "xA. b`(a`(x)) = x" +lemma comp_conv_id: assumes A1: "a bij(A,B)" and A2: "b:BA" and + A3: "xA. b`(a`(x)) = x" shows "b bij(B,A)" and "a = converse(b)" and "b = converse(a)" proof - have "b surj(B,A)" @@ -1569,50 +1582,50 @@

Theory func1

proof - { fix x assume "xA" let ?y = "a`(x)" - from A1 A3 xA have "?yB" and "b`(?y) = x" + from A1 A3 xA have "?yB" and "b`(?y) = x" using bij_def inj_def apply_funtype by auto hence "yB. b`(y) = x" by auto } thus ?thesis by simp qed - with A2 show "b surj(B,A)" using surj_def by simp + with A2 show "b surj(B,A)" using surj_def by simp qed moreover have "b inj(B,A)" proof - have "wB.yB. b`(w) = b`(y) w=y" proof - - { fix w y assume "wB" "yB" and I: "b`(w) = b`(y)" - from A1 have "a surj(A,B)" unfolding bij_def by simp - with wB obtain xw where "xw A" and II: "a`(xw) = w" + { fix w y assume "wB" "yB" and I: "b`(w) = b`(y)" + from A1 have "a surj(A,B)" unfolding bij_def by simp + with wB obtain xw where "xw A" and II: "a`(xw) = w" using surj_def by auto - with I have "b`(a`(xw)) = b`(y)" by simp + with I have "b`(a`(xw)) = b`(y)" by simp moreover from a surj(A,B) yB obtain xy where - "xy A" and III: "a`(xy) = y" + "xy A" and III: "a`(xy) = y" using surj_def by auto - moreover from A3 xw A xy A have "b`(a`(xw)) = xw" and "b`(a`(xy)) = xy" + moreover from A3 xw A xy A have "b`(a`(xw)) = xw" and "b`(a`(xy)) = xy" by auto ultimately have "xw = xy" by simp - with II III have "w=y" by simp + with II III have "w=y" by simp } thus ?thesis by auto qed - with A2 show "b inj(B,A)" using inj_def by auto + with A2 show "b inj(B,A)" using inj_def by auto qed ultimately show "b bij(B,A)" using bij_def by simp - from assms have "b O a = id(A)" using bij_def inj_def comp_eq_id_iff1 by auto - with A1 b bij(B,A) show "a = converse(b)" and "b = converse(a)" + from assms have "b O a = id(A)" using bij_def inj_def comp_eq_id_iff1 by auto + with A1 b bij(B,A) show "a = converse(b)" and "b = converse(a)" using comp_id_conv by auto qed text‹For a surjection the union if images of singletons is the whole range.› -lemma surj_singleton_image: assumes A1: "f surj(X,Y)" +lemma surj_singleton_image: assumes A1: "f surj(X,Y)" shows "(xX. {f`(x)}) = Y" proof - from A1 show "(xX. {f`(x)}) Y" + from A1 show "(xX. {f`(x)}) Y" using surj_def apply_funtype by auto next { fix y assume "y Y" - with A1 have "y (xX. {f`(x)})" + with A1 have "y (xX. {f`(x)})" using surj_def by auto } then show "Y (xX. {f`(x)})" by auto qed @@ -1629,56 +1642,56 @@

Theory func1

text‹We can create functions of two variables by combining functions of one variable.› -lemma cart_prod_fun: assumes "f1:X1Y1" "f2:X2Y2" and +lemma cart_prod_fun: assumes "f1:X1Y1" "f2:X2Y2" and "g = {p,f1`(fst(p)),f2`(snd(p)). p X1×X2}" - shows "g: X1×X2 Y1×Y2" using assms apply_funtype ZF_fun_from_total by simp + shows "g: X1×X2 Y1×Y2" using assms apply_funtype ZF_fun_from_total by simp text‹A reformulation of cart_prod_fun› above in a sligtly different notation.› -lemma prod_fun: +lemma prod_fun: assumes "f:X1X2" "g:X3X4" shows "{x,y,f`(x),g`(y). x,yX1×X3}:X1×X3X2×X4" proof - have "{x,y,f`(x),g`(y). x,yX1×X3} = {p,f`(fst(p)),g`(snd(p)). p X1×X3}" by auto - with assms show ?thesis using cart_prod_fun by simp + with assms show ?thesis using cart_prod_fun by simp qed text‹Product of two surjections is a surjection.› -theorem prod_functions_surj: +theorem prod_functions_surj: assumes "fsurj(A,B)" "gsurj(C,D)" shows "{a1,a2,f`(a1),g`(a2). a1,a2A×C} surj(A×C,B×D)" proof - let ?h = "{x, y, f`(x), g`(y) . x,y A×C}" - from assms have fun: "f:AB" "g:CD" unfolding surj_def by auto - then have pfun: "?h : A × C B × D" using prod_fun by auto + from assms have fun: "f:AB" "g:CD" unfolding surj_def by auto + then have pfun: "?h : A × C B × D" using prod_fun by auto { fix b assume "bB×D" then obtain b1 b2 where "b=b1,b2" "b1B" "b2D" by auto - with assms obtain a1 a2 where "f`(a1)=b1" "g`(a2)=b2" "a1A" "a2C" + with assms obtain a1 a2 where "f`(a1)=b1" "g`(a2)=b2" "a1A" "a2C" unfolding surj_def by blast hence "a1,a2,b1,b2 ?h" by auto - with pfun have "?h`a1,a2=b1,b2" using apply_equality by auto + with pfun have "?h`a1,a2=b1,b2" using apply_equality by auto with b=b1,b2 a1A a2C have "aA×C. ?h`(a)=b" by auto } hence "bB×D. aA×C. ?h`(a) = b" by auto - with pfun show ?thesis unfolding surj_def by auto + with pfun show ?thesis unfolding surj_def by auto qed text‹For a function of two variables created from functions of one variable as in cart_prod_fun› above, the inverse image of a cartesian product of sets is the cartesian product of inverse images.› -lemma cart_prod_fun_vimage: assumes "f1:X1Y1" "f2:X2Y2" and +lemma cart_prod_fun_vimage: assumes "f1:X1Y1" "f2:X2Y2" and "g = {p,f1`(fst(p)),f2`(snd(p)). p X1×X2}" shows "g-``(A1×A2) = f1-``(A1) × f2-``(A2)" proof - - from assms have "g: X1×X2 Y1×Y2" using cart_prod_fun + from assms have "g: X1×X2 Y1×Y2" using cart_prod_fun by simp then have "g-``(A1×A2) = {p X1×X2. g`(p) A1×A2}" using func1_1_L15 by simp - with assms g: X1×X2 Y1×Y2 show "g-``(A1×A2) = f1-``(A1) × f2-``(A2)" + with assms g: X1×X2 Y1×Y2 show "g-``(A1×A2) = f1-``(A1) × f2-``(A2)" using ZF_fun_from_tot_val func1_1_L15 by auto qed @@ -1701,20 +1714,20 @@

Theory func1

from the function. The next lemma is a technical fact that makes it easier to use this definition.›
-lemma fix_var_fun_domain: assumes A1: "f : X×Y Z" +lemma fix_var_fun_domain: assumes A1: "f : X×Y Z" shows "xX Fix1stVar(f,x) = {y,f`x,y. y Y}" "yY Fix2ndVar(f,y) = {x,f`x,y. x X}" proof - - from A1 have I: "domain(f) = X×Y" using func1_1_L1 by simp + from A1 have I: "domain(f) = X×Y" using func1_1_L1 by simp { assume "xX" - with I have "range(domain(f)) = Y" by auto + with I have "range(domain(f)) = Y" by auto then have "Fix1stVar(f,x) = {y,f`x,y. y Y}" using Fix1stVar_def by simp } then show "xX Fix1stVar(f,x) = {y,f`x,y. y Y}" by simp { assume "yY" - with I have "domain(domain(f)) = X" by auto + with I have "domain(domain(f)) = X" by auto then have "Fix2ndVar(f,y) = {x,f`x,y. x X}" using Fix2ndVar_def by simp } then show "yY Fix2ndVar(f,y) = {x,f`x,y. x X}" @@ -1723,26 +1736,26 @@

Theory func1

text‹If we fix the first variable, we get a function of the second variable.› -lemma fix_1st_var_fun: assumes A1: "f : X×Y Z" and A2: "xX" +lemma fix_1st_var_fun: assumes A1: "f : X×Y Z" and A2: "xX" shows "Fix1stVar(f,x) : Y Z" proof - - from A1 A2 have "yY. f`x,y Z" + from A1 A2 have "yY. f`x,y Z" using apply_funtype by simp then have "{y,f`x,y. y Y} : Y Z" using ZF_fun_from_total by simp - with A1 A2 show "Fix1stVar(f,x) : Y Z" using fix_var_fun_domain by simp + with A1 A2 show "Fix1stVar(f,x) : Y Z" using fix_var_fun_domain by simp qed text‹If we fix the second variable, we get a function of the first variable.› -lemma fix_2nd_var_fun: assumes A1: "f : X×Y Z" and A2: "yY" +lemma fix_2nd_var_fun: assumes A1: "f : X×Y Z" and A2: "yY" shows "Fix2ndVar(f,y) : X Z" proof - - from A1 A2 have "xX. f`x,y Z" + from A1 A2 have "xX. f`x,y Z" using apply_funtype by simp then have "{x,f`x,y. x X} : X Z" using ZF_fun_from_total by simp - with A1 A2 show "Fix2ndVar(f,y) : X Z" + with A1 A2 show "Fix2ndVar(f,y) : X Z" using fix_var_fun_domain by simp qed @@ -1750,26 +1763,26 @@

Theory func1

and the value of
Fix2ndVar(f,y)› at $x\in X$"?› lemma fix_var_val: - assumes A1: "f : X×Y Z" and A2: "xX" "yY" + assumes A1: "f : X×Y Z" and A2: "xX" "yY" shows "Fix1stVar(f,x)`(y) = f`x,y" "Fix2ndVar(f,y)`(x) = f`x,y" proof - let ?f1 = "{y,f`x,y. y Y}" let ?f2 = "{x,f`x,y. x X}" - from A1 A2 have I: + from A1 A2 have I: "Fix1stVar(f,x) = ?f1" "Fix2ndVar(f,y) = ?f2" using fix_var_fun_domain by auto - moreover from A1 A2 have + moreover from A1 A2 have "Fix1stVar(f,x) : Y Z" "Fix2ndVar(f,y) : X Z" using fix_1st_var_fun fix_2nd_var_fun by auto ultimately have "?f1 : Y Z" and "?f2 : X Z" by auto - with A2 have "?f1`(y) = f`x,y" and "?f2`(x) = f`x,y" + with A2 have "?f1`(y) = f`x,y" and "?f2`(x) = f`x,y" using ZF_fun_from_tot_val by auto - with I show + with I show "Fix1stVar(f,x)`(y) = f`x,y" "Fix2ndVar(f,y)`(x) = f`x,y" by auto @@ -1778,26 +1791,26 @@

Theory func1

text‹Fixing the second variable commutes with restrictig the domain.› lemma fix_2nd_var_restr_comm: - assumes A1: "f : X×Y Z" and A2: "yY" and A3: "X1 X" + assumes A1: "f : X×Y Z" and A2: "yY" and A3: "X1 X" shows "Fix2ndVar(restrict(f,X1×Y),y) = restrict(Fix2ndVar(f,y),X1)" proof - let ?g = "Fix2ndVar(restrict(f,X1×Y),y)" let ?h = "restrict(Fix2ndVar(f,y),X1)" - from A3 have I: "X1×Y X×Y" by auto - with A1 have II: "restrict(f,X1×Y) : X1×Y Z" + from A3 have I: "X1×Y X×Y" by auto + with A1 have II: "restrict(f,X1×Y) : X1×Y Z" using restrict_type2 by simp - with A2 have "?g : X1 Z" + with A2 have "?g : X1 Z" using fix_2nd_var_fun by simp moreover - from A1 A2 have III: "Fix2ndVar(f,y) : X Z" + from A1 A2 have III: "Fix2ndVar(f,y) : X Z" using fix_2nd_var_fun by simp - with A3 have "?h : X1 Z" + with A3 have "?h : X1 Z" using restrict_type2 by simp moreover - { fix z assume A4: "z X1" - with A2 I II have "?g`(z) = f`z,y" + { fix z assume A4: "z X1" + with A2 I II have "?g`(z) = f`z,y" using restrict fix_var_val by simp - also from A1 A2 A3 A4 have "f`z,y = ?h`(z)" + also from A1 A2 A3 A4 have "f`z,y = ?h`(z)" using restrict fix_var_val by auto finally have "?g`(z) = ?h`(z)" by simp } then have "z X1. ?g`(z) = ?h`(z)" by simp @@ -1807,25 +1820,25 @@

Theory func1

text‹The next lemma expresses the inverse image of a set by function with fixed first variable in terms of the original function.› -lemma fix_1st_var_vimage: +lemma fix_1st_var_vimage: assumes A1: "f : X×Y Z" and A2: "xX" shows "Fix1stVar(f,x)-``(A) = {yY. x,y f-``(A)}" proof - - from assms have "Fix1stVar(f,x)-``(A) = {yY. Fix1stVar(f,x)`(y) A}" + from assms have "Fix1stVar(f,x)-``(A) = {yY. Fix1stVar(f,x)`(y) A}" using fix_1st_var_fun func1_1_L15 by blast - with assms show ?thesis using fix_var_val func1_1_L15 by auto + with assms show ?thesis using fix_var_val func1_1_L15 by auto qed text‹The next lemma expresses the inverse image of a set by function with fixed second variable in terms of the original function.› -lemma fix_2nd_var_vimage: +lemma fix_2nd_var_vimage: assumes A1: "f : X×Y Z" and A2: "yY" shows "Fix2ndVar(f,y)-``(A) = {xX. x,y f-``(A)}" proof - - from assms have I: "Fix2ndVar(f,y)-``(A) = {xX. Fix2ndVar(f,y)`(x) A}" + from assms have I: "Fix2ndVar(f,y)-``(A) = {xX. Fix2ndVar(f,y)`(x) A}" using fix_2nd_var_fun func1_1_L15 by blast - with assms show ?thesis using fix_var_val func1_1_L15 by auto + with assms show ?thesis using fix_var_val func1_1_L15 by auto qed end diff --git a/docs/IsarMathLib/index.html b/docs/IsarMathLib/index.html index e653bce..63392a3 100644 --- a/docs/IsarMathLib/index.html +++ b/docs/IsarMathLib/index.html @@ -150,6 +150,14 @@

Theories

  • Topology_ZF_1
  • +
  • Topology_ZF_2
  • + +
  • Topology_ZF_4
  • + +
  • Topology_ZF_4a
  • + +
  • UniformSpace_ZF
  • +
  • MetricSpace_ZF
  • Real_ZF_2
  • @@ -162,16 +170,10 @@

    Theories

  • Topology_ZF_1b
  • -
  • Topology_ZF_2
  • -
  • Ring_Zariski_ZF_3
  • Topology_ZF_3
  • -
  • Topology_ZF_4
  • - -
  • Topology_ZF_4a
  • -
  • Topology_ZF_examples
  • Topology_ZF_examples_1
  • @@ -198,8 +200,6 @@

    Theories

  • MetricSpace_ZF_1
  • -
  • UniformSpace_ZF
  • -
  • UniformSpace_ZF_1
  • UniformSpace_ZF_2
  • diff --git a/docs/IsarMathLib/outline.pdf b/docs/IsarMathLib/outline.pdf index 71d5805..20c5285 100644 Binary files a/docs/IsarMathLib/outline.pdf and b/docs/IsarMathLib/outline.pdf differ diff --git a/docs/IsarMathLib/session_graph.pdf b/docs/IsarMathLib/session_graph.pdf index 0a02286..5dcf7fe 100644 Binary files a/docs/IsarMathLib/session_graph.pdf and b/docs/IsarMathLib/session_graph.pdf differ