Theory TopologicalGroup_ZF_1

(* 
    This file is a part of IsarMathLib - 
    a library of formalized mathematics written for Isabelle/Isar.

    Copyright (C) 2013 Daniel de la Concepcion

    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:

   1. Redistributions of source code must retain the above copyright notice, 
   this list of conditions and the following disclaimer.
   2. Redistributions in binary form must reproduce the above copyright notice, 
   this list of conditions and the following disclaimer in the documentation and/or 
   other materials provided with the distribution.
   3. The name of the author may not be used to endorse or promote products 
   derived from this software without specific prior written permission.

THIS SOFTWARE IS PROVIDED BY THE AUTHOR ``AS IS'' AND ANY EXPRESS OR IMPLIED 
WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE IMPLIED WARRANTIES OF 
MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE DISCLAIMED. 
IN NO EVENT SHALL THE AUTHOR BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, 
SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT LIMITED TO, 
PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, DATA, OR PROFITS; 
OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF LIABILITY, 
WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING NEGLIGENCE OR 
OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS SOFTWARE, 
EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. *)

section Topological groups 1

theory TopologicalGroup_ZF_1 imports TopologicalGroup_ZF Topology_ZF_properties_2
begin

textThis theory deals with some topological properties of topological groups.

subsectionSeparation properties of topological groups
textThe topological groups have very specific properties. For instance, $G$ is T0
iff it is T3.

theorem(in topgroup) cl_point:
  assumes "xG"
  shows "cl({x}) = (H𝒩0. x\<ltr>H)"
proof-
  {
    have c:"cl({x}) = (H𝒩0. {x}\<sad>H)" using cl_topgroup assms by auto
    {
      fix H
      assume "H𝒩0"
      then have "{x}\<sad>H=x\<ltr> H" using interval_add(3) assms
        by auto
      with H𝒩0 have "{x}\<sad>H{x\<ltr>H. H𝒩0}" by auto
    }
    then have "{{x}\<sad>H. H𝒩0}{x\<ltr>H. H𝒩0}" by auto
    moreover
    {
      fix H
      assume "H𝒩0"
      then have "{x}\<sad>H=x\<ltr> H" using interval_add(3) assms
        by auto
      with H𝒩0 have "x\<ltr> H{{x}\<sad>H. H𝒩0}" by auto
    }
    then have "{x\<ltr>H. H𝒩0}{{x}\<sad>H. H𝒩0}" by auto
    ultimately have "{{x}\<sad>H. H𝒩0}={x\<ltr>H. H𝒩0}" by auto
    then have "(H𝒩0. {x}\<sad>H) = (H𝒩0. x\<ltr>H)" by auto
    with c show "cl({x})=(H𝒩0. x\<ltr>H)" by auto
  }
qed

textWe prove the equivalence between $T_0$ and $T_1$ first.

theorem (in topgroup) neu_closed_imp_T1:
  assumes "{𝟬}{is closed in}T"
  shows "T{is T1}"
proof-
  {
    fix x z assume xG:"xG" and zG:"zG" and dis:"xz"
    then have clx:"cl({x})=(H𝒩0. x\<ltr>H)" using cl_point by auto
    {
      fix y
      assume "ycl({x})"
      with clx have "y(H𝒩0. x\<ltr>H)" by auto
      then have t:"H𝒩0. yx\<ltr>H" by auto
      from ycl({x}) xG have yG:"yG" using Top_3_L11(1) G_def by auto
      {
        fix H
        assume HNeig:"H𝒩0"
        with t have "yx\<ltr>H" by auto
        then obtain n where "y=x\<ra>n" and "nH" unfolding ltrans_def grop_def LeftTranslation_def by auto
        with HNeig have nG:"nG" unfolding zerohoods_def by auto
        from y=x\<ra>n and nH have "(\<rm>x)\<ra>yH" using group0.group0_2_L18(2) group0_valid_in_tgroup xG nG yG unfolding grinv_def grop_def
          by auto
      }
      then have el:"(\<rm>x)\<ra>y(𝒩0)" using zneigh_not_empty by auto
      have "cl({𝟬})=(H𝒩0. 𝟬\<ltr>H)" using cl_point zero_in_tgroup by auto
      moreover
      {
        fix H  assume "H𝒩0"
        then have "HG" unfolding zerohoods_def by auto
        then have "𝟬\<ltr>H=H" using image_id_same  group0.trans_neutral(2) group0_valid_in_tgroup unfolding gzero_def ltrans_def
          by auto
        with H𝒩0 have "𝟬\<ltr>H𝒩0" "H{𝟬\<ltr>H. H𝒩0}" by auto
      }
      then have "{𝟬\<ltr>H. H𝒩0}=𝒩0" by blast
      ultimately have "cl({𝟬})=(𝒩0)" by auto
      with el have "(\<rm>x)\<ra>ycl({𝟬})" by auto
      then have "(\<rm>x)\<ra>y{𝟬}" using assms Top_3_L8 G_def zero_in_tgroup by auto
      then have "(\<rm>x)\<ra>y=𝟬" by auto
      then have "y=\<rm>(\<rm>x)" using group0.group0_2_L9(2) group0_valid_in_tgroup neg_in_tgroup xG yG unfolding grop_def grinv_def by auto
      then have "y=x" using group0.group_inv_of_inv group0_valid_in_tgroup xG unfolding grinv_def by auto
    }
    then have "cl({x}){x}" by auto
    then have "cl({x})={x}" using xG cl_contains_set G_def by blast
    then have "{x}{is closed in}T" using Top_3_L8 xG G_def by auto
    then have "(T)-{x}T" using IsClosed_def by auto moreover
    from dis zG G_def have "z((T)-{x})  x((T)-{x})" by auto
    ultimately have "VT. zVxV" by(safe,auto)
  }
  then show "T{is T1}" using isT1_def by auto
qed

theorem (in topgroup) T0_imp_neu_closed:
  assumes "T{is T0}"
  shows "{𝟬}{is closed in}T"
proof-
  {
    fix x assume "xcl({𝟬})" and "x𝟬"
    have "cl({𝟬})=(H𝒩0. 𝟬\<ltr>H)" using cl_point zero_in_tgroup by auto
    moreover
    {
      fix H  assume "H𝒩0"
      then have "HG" unfolding zerohoods_def by auto
      then have "𝟬\<ltr>H=H" using image_id_same  group0.trans_neutral(2) group0_valid_in_tgroup unfolding gzero_def ltrans_def
        by auto
      with H𝒩0 have "𝟬\<ltr>H𝒩0" "H{𝟬\<ltr>H. H𝒩0}" by auto
    }
    then have "{𝟬\<ltr>H. H𝒩0}=𝒩0" by blast
    ultimately have "cl({𝟬})=(𝒩0)" by auto
    from x𝟬 and xcl({𝟬}) obtain U where "UT" and "(xU𝟬U)(𝟬UxU)" using assms Top_3_L11(1)
      zero_in_tgroup unfolding isT0_def G_def by blast moreover
    {
      assume "𝟬U"
      with UT have "U𝒩0" using zerohoods_def G_def Top_2_L3 by auto
      with xcl({𝟬}) and cl({𝟬})=(𝒩0) have "xU" by auto
    }
    ultimately have "𝟬U" and "xU" by auto
    with UT xcl({𝟬}) have "False" using cl_inter_neigh zero_in_tgroup unfolding G_def by blast
  }
  then have "cl({𝟬}){𝟬}" by auto
  then have "cl({𝟬})={𝟬}" using zero_in_tgroup cl_contains_set G_def by blast
  then show ?thesis using Top_3_L8 zero_in_tgroup unfolding G_def by auto
qed

subsectionExistence of nice neighbourhoods.


lemma (in topgroup) exist_basehoods_closed:
  assumes "U𝒩0"
  shows "V𝒩0. cl(V)U"
proof-
  from assms obtain V where "V𝒩0" "VU" "(V\<sad>V)U" "(\<sm>V)=V" using exists_procls_zerohood by blast
  have inv_fun:"GroupInv(G,f)GG" using group0_2_T2 Ggroup by auto
  have f_fun:"fG×GG" using group0.group_oper_fun group0_valid_in_tgroup by auto
  {
    fix x assume "xcl(V)"
    with V𝒩0 have "xT" "VT" using Top_3_L11(1) unfolding zerohoods_def G_def by blast+
    with V𝒩0 have "xint(x\<ltr>V)" using elem_in_int_ltrans G_def by auto
    with VTxcl(V) have "int(x\<ltr>V)V0" using cl_inter_neigh Top_2_L2 by blast
    then have "(x\<ltr>V)V0" using Top_2_L1 by blast
    then obtain q where "q(x\<ltr>V)" and "qV" by blast
    with VTxT obtain v where "q=x\<ra>v" "vV" unfolding ltrans_def grop_def using group0.ltrans_image
      group0_valid_in_tgroup unfolding G_def by auto
    from VT vVqV have "vT" "qT" by auto
    with q=x\<ra>vxT have "q\<rs>v=x" using group0.group0_2_L18(1) group0_valid_in_tgroup unfolding G_def
        unfolding grsub_def grinv_def grop_def by auto moreover
    from vV have "(\<rm>v)(\<sm>V)" unfolding setninv_def grinv_def using func_imagedef inv_fun VT G_def by auto
    then have "(\<rm>v)V" using (\<sm>V)=V by auto
    with qV have "q,\<rm>vV×V" by auto
    then have "f`q,\<rm>vV\<sad>V" using lift_subset_suff f_fun VT unfolding setadd_def by auto
    with V\<sad>VU have "q\<rs>vU" unfolding grsub_def grop_def by auto
    with q\<rs>v=x have "xU" by auto
  }
  then have "cl(V)U" by auto
  with V𝒩0 show ?thesis by auto
qed

subsectionRest of separation axioms

theorem(in topgroup) T1_imp_T2:
  assumes "T{is T1}"
  shows "T{is T2}"
proof-
  {
    fix x y assume ass:"xT" "yT" "xy"
    {
      assume "(\<rm>y)\<ra>x=𝟬"
      with ass(1,2) have "y=x" using group0.group0_2_L11[where a="y" and b="x"] group0_valid_in_tgroup by auto (*cannot be erased.*)
      with ass(3) have "False" by auto
    }
    then have "(\<rm>y)\<ra>x𝟬" by auto
    then have "𝟬(\<rm>y)\<ra>x" by auto
    from yT have "(\<rm>y)T" using neg_in_tgroup G_def by auto
    with xT have "(\<rm>y)\<ra>xT" using group0.group_op_closed[where a="\<rm>y" and b="x"] group0_valid_in_tgroup unfolding (*cannot be erased.*)
      G_def by auto
    with assms 𝟬(\<rm>y)\<ra>x obtain U where "UT" and "(\<rm>y)\<ra>xU" and "𝟬U" unfolding isT1_def using zero_in_tgroup
      by auto
    then have "U𝒩0" unfolding zerohoods_def G_def using Top_2_L3 by auto
    then obtain Q where "Q𝒩0" "QU" "(Q\<sad>Q)U" "(\<sm>Q)=Q" using exists_procls_zerohood by blast
    with (\<rm>y)\<ra>xU have "(\<rm>y)\<ra>xQ" by auto
    from Q𝒩0 have "QG" unfolding zerohoods_def by auto
    {
      assume "xy\<ltr>Q"
      with QG yT obtain u where "uQ" and "x=y\<ra>u" unfolding ltrans_def grop_def using group0.ltrans_image group0_valid_in_tgroup
        unfolding G_def by auto
      with QG have "uT" unfolding G_def by auto
      with x=y\<ra>u yT xT QG have "(\<rm>y)\<ra>x=u" using group0.group0_2_L18(2) group0_valid_in_tgroup unfolding G_def
        unfolding grsub_def grinv_def grop_def by auto
      with uQ have "(\<rm>y)\<ra>xQ" by auto
      then have "False" using (\<rm>y)\<ra>xQ by auto
    }
    then have "xy\<ltr>Q" by auto moreover
    {
      assume "yx\<ltr>Q"
      with QG xT obtain u where "uQ" and "y=x\<ra>u" unfolding ltrans_def grop_def using group0.ltrans_image group0_valid_in_tgroup
        unfolding G_def by auto
      with QG have "uT" unfolding G_def by auto
      with y=x\<ra>u yT xT QG have "(\<rm>x)\<ra>y=u" using group0.group0_2_L18(2) group0_valid_in_tgroup unfolding G_def
        unfolding grsub_def grinv_def grop_def by auto
      with uQ have "(\<rm>y)\<ra>x=\<rm>u" using group0.group_inv_of_two[OF group0_valid_in_tgroup group0.inverse_in_group[OF group0_valid_in_tgroup,of x],of y] (*From here no checked*)
        using xT yT using group0.group_inv_of_inv[OF group0_valid_in_tgroup] unfolding G_def grinv_def grop_def by auto
      moreover from uQ have "(\<rm>u)(\<sm>Q)" unfolding setninv_def grinv_def using func_imagedef[OF group0_2_T2[OF Ggroup] QG] by auto
      ultimately have "(\<rm>y)\<ra>xQ" using (\<rm>y)\<ra>xQ (\<sm>Q)=Q unfolding setninv_def grinv_def by auto
      then have "False" using (\<rm>y)\<ra>xQ by auto
    }
    then have "yx\<ltr>Q" by auto moreover
    {
      fix t
      assume "t(x\<ltr>Q)(y\<ltr>Q)"
      then have "t(x\<ltr>Q)" "t(y\<ltr>Q)" by auto
      with QG xT yT obtain u v where "uQ" "vQ" and "t=x\<ra>u" "t=y\<ra>v" unfolding ltrans_def grop_def using group0.ltrans_image[OF group0_valid_in_tgroup]
        unfolding G_def by auto
      then have "x\<ra>u=y\<ra>v" by auto
      moreover from uQ vQ QG have "uT" "vT" unfolding G_def by auto
      moreover note xT yT
      ultimately have "(\<rm>y)\<ra>(x\<ra>u)=v" using group0.group0_2_L18(2)[OF group0_valid_in_tgroup, of y v "x\<ra>u"] group0.group_op_closed[OF group0_valid_in_tgroup, of x u] unfolding G_def
        unfolding grsub_def grinv_def grop_def by auto
      then have "((\<rm>y)\<ra>x)\<ra>u=v" using group0.group_oper_assoc[OF group0_valid_in_tgroup]
        unfolding grop_def using xT yT uT using group0.inverse_in_group[OF group0_valid_in_tgroup] unfolding G_def
        by auto
      then have "((\<rm>y)\<ra>x)=v\<rs>u" using group0.group0_2_L18(1)[OF group0_valid_in_tgroup,of "(\<rm>y)\<ra>x" u v]
        using (\<rm>y)\<ra>xT uT vT unfolding G_def grsub_def grinv_def grop_def by force
      moreover 
      from uQ have "(\<rm>u)(\<sm>Q)" unfolding setninv_def grinv_def using func_imagedef[OF group0_2_T2[OF Ggroup] QG] by auto
      then have "(\<rm>u)Q" using (\<sm>Q)=Q by auto
      with vQ have "v,\<rm>uQ×Q" by auto
      then have "f`v,\<rm>uQ\<sad>Q" using lift_subset_suff[OF group0.group_oper_fun[OF group0_valid_in_tgroup] QG QG]
        unfolding setadd_def by auto
      with Q\<sad>QU have "v\<rs>uU" unfolding grsub_def grop_def by auto
      ultimately have "(\<rm>y)\<ra>xU" by auto
      with (\<rm>y)\<ra>xU have "False" by auto
    }
    then have "(x\<ltr>Q)(y\<ltr>Q)=0" by auto
    moreover have "xint(x\<ltr>Q)""yint(y\<ltr>Q)" using elem_in_int_ltrans Q𝒩0
      xT yT unfolding G_def by auto moreover
    have "int(x\<ltr>Q)(x\<ltr>Q)""int(y\<ltr>Q)(y\<ltr>Q)" using Top_2_L1 by auto
    moreover have "int(x\<ltr>Q)T" "int(y\<ltr>Q)T" using Top_2_L2 by auto
    ultimately have "int(x\<ltr>Q)T  int(y\<ltr>Q)T  x  int(x\<ltr>Q)  y  int(y\<ltr>Q)  int(x\<ltr>Q)  int(y\<ltr>Q) = 0"
      by blast
    then have "UT. VT. xUyVUV=0" by auto
  }
  then show ?thesis using isT2_def by auto
qed

textHere follow some auxiliary lemmas.

lemma (in topgroup) trans_closure:
  assumes "xG" "AG"
  shows "cl(x\<ltr>A)=x\<ltr>cl(A)"
proof-
  have "T-(T-(x\<ltr>A))=(x\<ltr>A)" unfolding ltrans_def using group0.group0_5_L1(2)[OF group0_valid_in_tgroup assms(1)]
    unfolding image_def range_def domain_def converse_def Pi_def by auto
  then have "cl(x\<ltr>A)=T-int(T-(x\<ltr>A))" using Top_3_L11(2)[of "T-(x\<ltr>A)"] by auto moreover
  have "x\<ltr>G=G" using surj_image_eq group0.trans_bij(2)[OF group0_valid_in_tgroup assms(1)] bij_def by auto
  then have "T-(x\<ltr>A)=x\<ltr>(T-A)" using inj_image_dif[of "LeftTranslation(G, f, x)""G""G", OF _ assms(2)]
    unfolding ltrans_def G_def using group0.trans_bij(2)[OF group0_valid_in_tgroup assms(1)] bij_def by auto
  then have "int(T-(x\<ltr>A))=int(x\<ltr>(T-A))" by auto
  then have "int(T-(x\<ltr>A))=x\<ltr>int(T-A)" using ltrans_interior[OF assms(1),of "T-A"] unfolding G_def by force
  have "T-int(T-A)=cl(T-(T-A))" using Top_3_L11(2)[of "T-A"] by force
  have "T-(T-A)=A" using assms(2) G_def by auto
  with T-int(T-A)=cl(T-(T-A)) have "T-int(T-A)=cl(A)" by auto
  have "T-(T-int(T-A))=int(T-A)" using Top_2_L2 by auto
  with T-int(T-A)=cl(A) have "int(T-A)=T-cl(A)" by auto
  with int(T-(x\<ltr>A))=x\<ltr>int(T-A) have "int(T-(x\<ltr>A))=x\<ltr>(T-cl(A))" by auto
  with x\<ltr>G=G have "int(T-(x\<ltr>A))=T-(x\<ltr>cl(A))" using inj_image_dif[of "LeftTranslation(G, f, x)""G""G""cl(A)"]
    unfolding ltrans_def using group0.trans_bij(2)[OF group0_valid_in_tgroup assms(1)] Top_3_L11(1) assms(2) unfolding bij_def G_def
    by auto
  then have "T-int(T-(x\<ltr>A))=T-(T-(x\<ltr>cl(A)))" by auto
  then have "T-int(T-(x\<ltr>A))=x\<ltr>cl(A)" unfolding ltrans_def using group0.group0_5_L1(2)[OF group0_valid_in_tgroup assms(1)]
    unfolding image_def range_def domain_def converse_def Pi_def by auto
  with cl(x\<ltr>A)=T-int(T-(x\<ltr>A)) show ?thesis by auto
qed

lemma (in topgroup) trans_interior2: assumes A1: "gG" and A2: "AG" 
  shows "int(A)\<rtr>g = int(A\<rtr>g)"
proof -
  from assms have "A  T" and "IsAhomeomorphism(T,T,RightTranslation(G,f,g))"
    using tr_homeo by auto
  then show ?thesis using int_top_invariant by simp
qed

lemma (in topgroup) trans_closure2:
  assumes "xG" "AG"
  shows "cl(A\<rtr>x)=cl(A)\<rtr>x"
proof-
  have "T-(T-(A\<rtr>x))=(A\<rtr>x)" unfolding ltrans_def using group0.group0_5_L1(1)[OF group0_valid_in_tgroup assms(1)]
    unfolding image_def range_def domain_def converse_def Pi_def by auto
  then have "cl(A\<rtr>x)=T-int(T-(A\<rtr>x))" using Top_3_L11(2)[of "T-(A\<rtr>x)"] by auto moreover
  have "G\<rtr>x=G" using surj_image_eq group0.trans_bij(1)[OF group0_valid_in_tgroup assms(1)] bij_def by auto
  then have "T-(A\<rtr>x)=(T-A)\<rtr>x" using inj_image_dif[of "RightTranslation(G, f, x)""G""G", OF _ assms(2)]
    unfolding rtrans_def G_def using group0.trans_bij(1)[OF group0_valid_in_tgroup assms(1)] bij_def by auto
  then have "int(T-(A\<rtr>x))=int((T-A)\<rtr>x)" by auto
  then have "int(T-(A\<rtr>x))=int(T-A)\<rtr>x" using trans_interior2[OF assms(1),of "T-A"] unfolding G_def by force
  have "T-int(T-A)=cl(T-(T-A))" using Top_3_L11(2)[of "T-A"] by force
  have "T-(T-A)=A" using assms(2) G_def by auto
  with T-int(T-A)=cl(T-(T-A)) have "T-int(T-A)=cl(A)" by auto
  have "T-(T-int(T-A))=int(T-A)" using Top_2_L2 by auto
  with T-int(T-A)=cl(A) have "int(T-A)=T-cl(A)" by auto
  with int(T-(A\<rtr>x))=int(T-A)\<rtr>x have "int(T-(A\<rtr>x))=(T-cl(A))\<rtr>x" by auto
  with G\<rtr>x=G have "int(T-(A\<rtr>x))=T-(cl(A)\<rtr>x)" using inj_image_dif[of "RightTranslation(G, f, x)""G""G""cl(A)"]
    unfolding rtrans_def using group0.trans_bij(1)[OF group0_valid_in_tgroup assms(1)] Top_3_L11(1) assms(2) unfolding bij_def G_def
    by auto
  then have "T-int(T-(A\<rtr>x))=T-(T-(cl(A)\<rtr>x))" by auto
  then have "T-int(T-(A\<rtr>x))=cl(A)\<rtr>x" unfolding ltrans_def using group0.group0_5_L1(1)[OF group0_valid_in_tgroup assms(1)]
    unfolding image_def range_def domain_def converse_def Pi_def by auto
  with cl(A\<rtr>x)=T-int(T-(A\<rtr>x)) show ?thesis by auto
qed

lemma (in topgroup) trans_subset:
  assumes "A((\<rm>x)\<ltr>B)""xG""AG""BG"
  shows "x\<ltr>AB"
proof-
  {
   fix t assume "tx\<ltr>A"
    with xG AG obtain u where "uA" "t=x\<ra>u" unfolding ltrans_def grop_def using group0.ltrans_image[OF group0_valid_in_tgroup]
      unfolding G_def by auto
    with xG AG uA have "(\<rm>x)\<ra>t=u" using group0.group0_2_L18(2)[OF group0_valid_in_tgroup, of "x""u""t"]
      group0.group_op_closed[OF group0_valid_in_tgroup,of x u] unfolding grop_def grinv_def by auto
    with uA have "(\<rm>x)\<ra>tA" by auto
    with A(\<rm>x)\<ltr>B have "(\<rm>x)\<ra>t(\<rm>x)\<ltr>B" by auto
    with BG obtain v where "(\<rm>x)\<ra>t=(\<rm>x)\<ra>v" "vB" unfolding ltrans_def grop_def using neg_in_tgroup[OF xG] group0.ltrans_image[OF group0_valid_in_tgroup]
      unfolding G_def by auto
    have "LeftTranslation(G,f,\<rm>x)inj(G,G)" using group0.trans_bij(2)[OF group0_valid_in_tgroup neg_in_tgroup[OF xG]] bij_def by auto
    then have eq:"AG. BG. LeftTranslation(G,f,\<rm>x)`A=LeftTranslation(G,f,\<rm>x)`B  A=B" unfolding inj_def by auto
    {
      fix A B assume "AG""BG"
      assume "f`\<rm>x,A=f`\<rm>x,B"
      then have "LeftTranslation(G,f,\<rm>x)`A=LeftTranslation(G,f,\<rm>x)`B" using group0.group0_5_L2(2)[OF group0_valid_in_tgroup neg_in_tgroup[OF xG]]
        AGBG by auto
      with eq AGBG have "A=B" by auto
    }
    then have eq1:"AG. BG. f`\<rm>x,A=f`\<rm>x,B  A=B" by auto
    from AG uA have "uG" by auto
    with vB BG t=x\<ra>u have "tG" "vG" using group0.group_op_closed[OF group0_valid_in_tgroup xG,of u] unfolding grop_def
      by auto
    with eq1 (\<rm>x)\<ra>t=(\<rm>x)\<ra>v have "t=v" unfolding grop_def by auto
    with vB have "tB" by auto
  }
  then show ?thesis by auto
qed

textEvery topological group is regular, and hence $T_3$. The proof is in the next
section, since it uses local properties.

subsectionLocal properties

textIn a topological group, all local properties depend only on the neighbourhoods
of the neutral element; when considering topological properties. The next result
of regularity, will use this idea, since translations preserve closed sets.

lemma (in topgroup) local_iff_neutral:
  assumes "UT𝒩0. N𝒩0. NU P(N,T)" "NPow(G). xG. P(N,T)  P(x\<ltr>N,T)"
  shows "T{is locally}P"
proof-
  {
    fix x U assume "xT""UT""xU"
    then have "(\<rm>x)\<ltr>UT𝒩0" using open_tr_open(1) open_trans_neigh neg_in_tgroup unfolding G_def
      by auto
    with assms(1) obtain N where "N((\<rm>x)\<ltr>U)""P(N,T)""N𝒩0" by auto
    note xTN((\<rm>x)\<ltr>U) moreover
    from UT have "UT" by auto moreover
    from N𝒩0 have "NG" unfolding zerohoods_def by auto
    ultimately have "(x\<ltr>N)U" using trans_subset unfolding G_def by auto moreover
    from NGxT assms(2) P(N,T) have "P((x\<ltr>N),T)" unfolding G_def by auto moreover
    from N𝒩0xT have "xint(x\<ltr>N)" using elem_in_int_ltrans unfolding G_def by auto
    ultimately have "NPow(U). xint(N)P(N,T)" by auto
  }
  then show ?thesis unfolding IsLocally_def[OF topSpaceAssum] by auto
qed

lemma (in topgroup) trans_closed:
  assumes "A{is closed in}T""xG"
  shows "(x\<ltr>A){is closed in}T"
proof-
  from assms(1) have "cl(A)=A" using Top_3_L8 unfolding IsClosed_def by auto
  then have "x\<ltr>cl(A)=x\<ltr>A" by auto
  then have "cl(x\<ltr>A)=x\<ltr>A" using trans_closure assms unfolding IsClosed_def by auto
  moreover have "x\<ltr>AG" unfolding ltrans_def using group0.group0_5_L1(2)[OF group0_valid_in_tgroup xG]
      unfolding image_def range_def domain_def converse_def Pi_def by auto
  ultimately show ?thesis using Top_3_L8 unfolding G_def by auto
qed

textAs it is written in the previous section, every topological group is regular.

theorem (in topgroup) topgroup_reg:
  shows "T{is regular}"
proof-
  {
    fix U assume "UT𝒩0"
    then obtain V where "cl(V)U""V𝒩0" using exist_basehoods_closed by blast
    then have "Vcl(V)" using cl_contains_set unfolding zerohoods_def G_def by auto
    then have "int(V)int(cl(V))" using interior_mono by auto
    with V𝒩0 have "cl(V)𝒩0" unfolding zerohoods_def G_def using Top_3_L11(1) by auto
    from V𝒩0 have "cl(V){is closed in}T" using cl_is_closed unfolding zerohoods_def G_def by auto
    with cl(V)𝒩0cl(V)U have "N𝒩0. NUN{is closed in}T" by auto
  }
  then have "UT𝒩0. N𝒩0. NUN{is closed in}T" by auto moreover
  have "NPow(G).( xG. (N{is closed in}T(x\<ltr>N){is closed in}T))" using trans_closed by auto
  ultimately have "T{is locally-closed}" using local_iff_neutral unfolding IsLocallyClosed_def by auto
  then show "T{is regular}" using regular_locally_closed by auto
qed

textThe promised corollary follows:
    
corollary (in topgroup) T2_imp_T3:
  assumes "T{is T2}"
  shows "T{is T3}" using T2_is_T1 topgroup_reg isT3_def assms by auto

end