Theory EquivClass1

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

    Copyright (C) 2005-2008  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:

   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 Equivalence relations

theory EquivClass1 imports ZF.EquivClass func_ZF ZF1

begin

textIn this theory file we extend the  work on equivalence relations 
  done in the standard Isabelle's EquivClass theory. That development
  is very good and all, but we really would prefer an approach contained within
  the a standard ZF set theory, without extensions specific to Isabelle.
  That is why this theory is written.


subsectionCongruent functions and projections on the quotient

textSuppose we have a set $X$ with a relation 
  $r\subseteq X\times X$ and a function $f: X\rightarrow X$. 
  The function $f$ can be compatible (congruent) with $r$ in the sense that
  if two elements $x,y$ are related then the values $f(x), f(x)$ 
  are also related. This is especially useful if $r$ is an 
  equivalence relation as it allows to "project" the function
  to the quotient space $X/r$ (the set of equivalence classes of $r$)
  and create a new function $F$ that satifies the formula
  $F([x]_r) = [f(x)]_r$. When $f$ is congruent with respect 
  to $r$ such definition of the value of $F$ on the equivalence class
  $[x]_r$ does not depend on which $x$ we choose to represent the class.
  In this section we also consider binary operations that are congruent
  with respect to a relation. These are important in algebra - the
  congruency condition allows to project the operation 
  to obtain the operation on the quotient space. 


textFirst we define the notion of function that maps equivalent 
  elements to equivalent values. We use similar names as
  in the Isabelle's standard EquivClass› theory to indicate 
  the conceptual correspondence of the notions.

definition
  "Congruent(r,f) 
  (x y. x,y  r   f`(x),f`(y)  r)"

textNow we will define the projection of
  a function onto the quotient space. In standard math the equivalence class
  of $x$ with respect to relation $r$ is usually denoted $[x]_r$. Here we reuse
  notation $r\{ x\}$ instead. This means the image of the set $\{ x\}$ 
  with respect to the relation, which, for equivalence relations is 
  exactly its equivalence class if you think about it.

definition
  "ProjFun(A,r,f) 
  {c,xc. r``{f`(x)}. c  (A//r)}"

textElements of equivalence classes belong to the set.

lemma EquivClass_1_L1: 
  assumes A1: "equiv(A,r)" and A2: "C  A//r" and A3: "xC"
  shows "xA"
proof -
  from A2 have "C   (A//r)" by auto
  with A1 A3 show "xA"
    using Union_quotient by auto
qed

textThe image of a subset of $X$ under projection is a subset
  of $A/r$.

lemma EquivClass_1_L1A: 
  assumes "AX" shows "{r``{x}. xA}  X//r"
  using assms quotientI by auto

textIf an element belongs to an equivalence class, then its image
  under relation is this equivalence class.

lemma EquivClass_1_L2: 
  assumes A1: "equiv(A,r)"  "C  A//r" and A2: "xC"
  shows "r``{x} = C"
proof -
  from A1 A2 have "x  r``{x}" 
    using EquivClass_1_L1  equiv_class_self by simp
  with A2 have I: "r``{x}C  0" by auto
  from A1 A2 have "r``{x}  A//r"
    using EquivClass_1_L1 quotientI by simp
  with A1 I show ?thesis
    using quotient_disj by blast
qed
      
textElements that belong to the same equivalence class are equivalent.

lemma EquivClass_1_L2A:
  assumes "equiv(A,r)"  "C  A//r"  "xC"  "yC"
  shows "x,y  r" 
  using assms EquivClass_1_L2 EquivClass_1_L1 equiv_class_eq_iff
  by simp

textElements that have the same image under an equivalence relation are equivalent.
  This is the same as eq_equiv_class› from standard Isabelle/ZF's EquivClass› theory, 
  just copied here to be easier to find.

lemma same_image_equiv: 
  assumes "equiv(A,r)" "yA" "r``{x} = r``{y}"
  shows "x,y  r" using assms eq_equiv_class by simp

textEvery $x$ is in the class of $y$, then they are equivalent.

lemma EquivClass_1_L2B: 
  assumes A1: "equiv(A,r)" and A2: "yA" and A3: "x  r``{y}"
  shows "x,y  r"
proof -
  from A2 have  "r``{y}  A//r"
    using quotientI by simp
  with A1 A3 show ?thesis using
    EquivClass_1_L1 equiv_class_self equiv_class_nondisjoint by blast
qed
  
textIf a function is congruent then the equivalence classes of the values
  that come from the arguments from the same class are the same.

lemma EquivClass_1_L3: 
  assumes A1: "equiv(A,r)" and A2: "Congruent(r,f)" 
  and A3: "C  A//r"  "xC"  "yC" 
  shows "r``{f`(x)} = r``{f`(y)}"
proof -
  from A1 A3 have "x,y  r"
    using EquivClass_1_L2A by simp
  with A2 have  "f`(x),f`(y)  r"
    using Congruent_def by simp
  with A1 show ?thesis using equiv_class_eq by simp
qed

textThe values of congruent functions are in the space.

lemma EquivClass_1_L4:
  assumes A1: "equiv(A,r)" and A2: "C  A//r"  "xC"
  and A3: "Congruent(r,f)"
  shows "f`(x)  A"
proof -
  from A1 A2 have "xA"
    using EquivClass_1_L1 by simp
  with A1 have "x,x  r"
    using equiv_def refl_def by simp
  with A3 have  "f`(x),f`(x)  r"
    using Congruent_def by simp
  with A1 show ?thesis using equiv_type by auto
qed

textEquivalence classes are not empty.

lemma EquivClass_1_L5: 
  assumes A1: "refl(A,r)" and A2: "C  A//r"
  shows "C0"
proof -
  from A2 obtain x where I: "C = r``{x}" and "xA"
    using quotient_def by auto
  from A1 xA have "x  r``{x}" using refl_def by auto
  with I show ?thesis by auto
qed
  
textTo avoid using an axiom of choice, we define the projection using 
  the expression $\bigcup _{x\in C} r(\{f(x)\})$. 
  The next lemma shows that for
  congruent function this is in the quotient space $A/r$.

lemma EquivClass_1_L6:
  assumes A1: "equiv(A,r)" and A2: "Congruent(r,f)" 
  and A3: "C  A//r"
  shows "(xC. r``{f`(x)})  A//r"
proof -
  from A1 have "refl(A,r)" unfolding equiv_def by simp
  with A3 have "C0" using EquivClass_1_L5 by simp
  moreover from A2 A3 A1 have "xC. r``{f`(x)}  A//r"
    using EquivClass_1_L4 quotientI by auto
  moreover from A1 A2 A3 have 
    "x y. xC  yC  r``{f`(x)} = r``{f`(y)}" 
    using EquivClass_1_L3 by blast
  ultimately show ?thesis by (rule ZF1_1_L2)
qed

textCongruent functions can be projected.

lemma EquivClass_1_T0: 
  assumes "equiv(A,r)"  "Congruent(r,f)"
  shows "ProjFun(A,r,f) : A//r  A//r"
  using assms EquivClass_1_L6 ProjFun_def ZF_fun_from_total
  by simp
  
textWe now define congruent functions of two variables (binary funtions). 
  The predicate Congruent2› corresponds to congruent2› 
  in Isabelle's standard EquivClass› theory, 
  but uses ZF-functions rather than meta-functions.

definition
  "Congruent2(r,f) 
  (x1 x2 y1 y2. x1,x2  r  y1,y2  r   
  f`x1,y1, f`x2,y2   r)"


textNext we define the notion of projecting a binary operation
  to the quotient space. This is a very important concept that
  allows to define quotient groups, among other things.

definition 
  "ProjFun2(A,r,f) 
  {p, z  fst(p)×snd(p). r``{f`(z)}. p  (A//r)×(A//r) }"


textThe following lemma is a two-variables equivalent 
  of EquivClass_1_L3›.

lemma EquivClass_1_L7: 
  assumes A1: "equiv(A,r)" and A2: "Congruent2(r,f)"
  and A3: "C1  A//r"  "C2  A//r" 
  and A4: "z1  C1×C2"  "z2  C1×C2"
  shows "r``{f`(z1)} = r``{f`(z2)}"
proof -
  from A4 obtain x1 y1 x2 y2 where 
    "x1C1" and "y1C2" and "z1 = x1,y1" and 
    "x2C1" and "y2C2" and "z2 = x2,y2" 
    by auto
  with A1 A3 have "x1,x2  r" and "y1,y2  r"
    using EquivClass_1_L2A by auto
  with A2 have "f`x1,y1,f`x2,y2  r"
    using Congruent2_def by simp
  with A1 z1 = x1,y1 z2 = x2,y2 show ?thesis 
    using equiv_class_eq by simp
qed

textThe values of congruent functions of two variables are in the space.

lemma EquivClass_1_L8:
  assumes A1: "equiv(A,r)" and A2: "C1  A//r" and A3: "C2  A//r"
  and A4: "z  C1×C2" and A5: "Congruent2(r,f)"
  shows "f`(z)  A"
proof -
  from A4 obtain x y where "xC1" and "yC2" and "z = x,y"  
    by auto
  with A1 A2 A3 have "xA" and "yA" 
    using EquivClass_1_L1 by auto
  with A1 A4 have "x,x  r" and "y,y  r"
    using equiv_def refl_def by auto
  with A5 have "f`x,y, f`x,y   r"
    using Congruent2_def by simp
  with A1 z = x,y show ?thesis using equiv_type by auto
qed

textThe values of congruent functions are in the space. Note that
  although this lemma is intended to be used with functions, we don't
  need to assume that $f$ is a function.

lemma EquivClass_1_L8A:
  assumes A1: "equiv(A,r)" and A2: "xA"  "yA"
  and A3: "Congruent2(r,f)"
  shows "f`x,y  A"
proof -
  from A1 A2 have "r``{x}  A//r" "r``{y}  A//r" 
    "x,y  r``{x}×r``{y}"
    using equiv_class_self quotientI by auto
  with A1 A3 show ?thesis using EquivClass_1_L8 by simp
qed
  
textThe following lemma is a two-variables equivalent of 
  EquivClass_1_L6›.

lemma EquivClass_1_L9:
  assumes A1: "equiv(A,r)" and A2: "Congruent2(r,f)" 
  and A3: "p  (A//r)×(A//r)"
  shows "( z  fst(p)×snd(p). r``{f`(z)})  A//r"
proof -
  from A3 have "fst(p)  A//r" and "snd(p)  A//r"
    by auto
  with A1 A2 have 
    I: "z  fst(p)×snd(p). f`(z)  A"
    using EquivClass_1_L8 by simp
  from A3 A1 have "fst(p)×snd(p)  0" 
    using equiv_def EquivClass_1_L5 Sigma_empty_iff
    by auto
  moreover from A1 I have 
    "z  fst(p)×snd(p). r``{f`(z)}  A//r"
    using quotientI by simp
  moreover from A1 A2 fst(p)  A//r snd(p)  A//r have
    "z1 z2. z1  fst(p)×snd(p)  z2  fst(p)×snd(p)  
    r``{f`(z1)} = r``{f`(z2)}"
    using EquivClass_1_L7 by blast
   ultimately show ?thesis by (rule ZF1_1_L2)
qed

textCongruent functions of two variables can be projected.

theorem EquivClass_1_T1: 
  assumes "equiv(A,r)"  "Congruent2(r,f)"
  shows "ProjFun2(A,r,f) : (A//r)×(A//r)  A//r"
  using assms EquivClass_1_L9 ProjFun2_def ZF_fun_from_total 
  by simp

textThe projection diagram commutes. I wish I knew how to draw this diagram
  in LaTeX.

lemma EquivClass_1_L10: 
  assumes A1: "equiv(A,r)" and A2: "Congruent2(r,f)" 
  and A3: "xA"  "yA"
  shows "ProjFun2(A,r,f)`r``{x},r``{y} = r``{f`x,y}"
proof -
  from A3 A1 have "r``{x} × r``{y}  0"
    using quotientI equiv_def EquivClass_1_L5 Sigma_empty_iff
    by auto
  moreover have 
    "z  r``{x}×r``{y}.  r``{f`(z)} = r``{f`x,y}"
  proof
    fix z assume A4: "z  r``{x}×r``{y}"
    from A1 A3 have 
      "r``{x}  A//r" "r``{y}  A//r"
      "x,y  r``{x}×r``{y}"
      using quotientI equiv_class_self by auto
    with A1 A2 A4 show
      "r``{f`(z)} = r``{f`x,y}"
      using EquivClass_1_L7 by blast
  qed
  ultimately have 
    "(z  r``{x}×r``{y}. r``{f`(z)}) =  r``{f`x,y}"
    by (rule ZF1_1_L1)
  moreover have 
    "ProjFun2(A,r,f)`r``{x},r``{y} = (z  r``{x}×r``{y}. r``{f`(z)})"
    proof -
      from assms have 
	"ProjFun2(A,r,f) : (A//r)×(A//r)  A//r"
	"r``{x},r``{y}  (A//r)×(A//r)"
	using EquivClass_1_T1 quotientI by auto
      then show ?thesis using ProjFun2_def ZF_fun_from_tot_val
	by auto
    qed
  ultimately show ?thesis by simp
qed

subsectionProjecting commutative, associative and distributive operations.

textIn this section we show that if the operations are congruent with
  respect to an equivalence relation then the projection to the quotient 
  space preserves commutativity, associativity and distributivity.

textThe projection of commutative operation is commutative.

lemma EquivClass_2_L1: assumes 
  A1: "equiv(A,r)" and A2: "Congruent2(r,f)"
  and A3: "f {is commutative on} A"
  and A4: "c1  A//r"  "c2  A//r"
  shows "ProjFun2(A,r,f)`c1,c2 = ProjFun2(A,r,f)`c2,c1"
proof -
  from A4 obtain x y where D1:
    "c1 = r``{x}"  "c2 = r``{y}"
    "xA"  "yA"
    using quotient_def by auto
  with A1 A2 have "ProjFun2(A,r,f)`c1,c2 = r``{f`x,y}"
    using EquivClass_1_L10 by simp
  also from A3 D1 have
    "r``{f`x,y} = r``{f`y,x}"
    using IsCommutative_def by simp
  also from A1 A2 D1 have
    "r``{f`y,x} = ProjFun2(A,r,f)` c2,c1"
    using EquivClass_1_L10 by simp
  finally show ?thesis by simp
qed

textThe projection of commutative operation is commutative.

theorem EquivClass_2_T1:
  assumes "equiv(A,r)" and "Congruent2(r,f)"
  and "f {is commutative on} A"
  shows "ProjFun2(A,r,f) {is commutative on} A//r"
  using assms IsCommutative_def EquivClass_2_L1 by simp
    
textThe projection of an associative operation is associative.

lemma EquivClass_2_L2: 
  assumes A1: "equiv(A,r)" and A2: "Congruent2(r,f)"
  and A3: "f {is associative on} A"
  and A4: "c1  A//r"  "c2  A//r"  "c3  A//r"
  and A5: "g = ProjFun2(A,r,f)"
  shows "g`g`c1,c2,c3 = g`c1,g`c2,c3"
proof -
  from A4 obtain x y z where D1:
    "c1 = r``{x}"  "c2 = r``{y}"  "c3 = r``{z}"
    "xA"  "yA"  "zA"
    using quotient_def by auto
  with A3 have T1:"f`x,y  A"  "f`y,z  A"
    using IsAssociative_def apply_type by auto
  with A1 A2 D1 A5 have 
    "g`g`c1,c2,c3 =  r``{f`f`x,y,z}"
    using EquivClass_1_L10 by simp
  also from D1 A3 have 
    " = r``{f`x,f`y,z }"
    using IsAssociative_def by simp
  also from T1 A1 A2 D1 A5 have
    " = g`c1,g`c2,c3"
    using EquivClass_1_L10 by simp
  finally show ?thesis by simp
qed

textThe projection of an associative operation is associative on the
  quotient.

theorem EquivClass_2_T2:
  assumes A1: "equiv(A,r)" and A2: "Congruent2(r,f)"
  and A3: "f {is associative on} A"
  shows "ProjFun2(A,r,f) {is associative on} A//r"
proof -
  let ?g = "ProjFun2(A,r,f)"
  from A1 A2 have 
    "?g  (A//r)×(A//r)  A//r"
    using EquivClass_1_T1 by simp
  moreover from A1 A2 A3 have
    "c1  A//r.c2  A//r.c3  A//r.
    ?g`?g`c1,c2,c3 = ?g`c1,?g`c2,c3"
    using EquivClass_2_L2 by simp
  ultimately show ?thesis
    using IsAssociative_def by simp
qed

textThe essential condition to show that distributivity is preserved by 
  projections to quotient spaces, provided both operations are congruent
  with respect to the equivalence relation.

lemma EquivClass_2_L3: 
  assumes A1: "IsDistributive(X,A,M)"
  and A2: "equiv(X,r)" 
  and A3: "Congruent2(r,A)" "Congruent2(r,M)"
  and A4: "a  X//r"  "b  X//r"  "c  X//r"
  and A5: "Ap = ProjFun2(X,r,A)" "Mp = ProjFun2(X,r,M)"
  shows "Mp`a,Ap`b,c = Ap` Mp`a,b,Mp`a,c  
  Mp` Ap`b,c,a  = Ap` Mp`b,a, Mp`c,a"
proof
  from A4 obtain x y z where "xX"  "yX"  "zX"
    "a = r``{x}"  "b = r``{y}"  "c = r``{z}"   
    using quotient_def by auto
  with A1 A2 A3 A5 show 
    "Mp`a,Ap`b,c = Ap` Mp`a,b,Mp`a,c" and
    "Mp` Ap`b,c,a  = Ap` Mp`b,a, Mp`c,a"  
    using EquivClass_1_L8A EquivClass_1_L10 IsDistributive_def
    by auto
qed

textDistributivity is preserved by 
  projections to quotient spaces, provided both operations are congruent
  with respect to the equivalence relation.

lemma EquivClass_2_L4: assumes A1: "IsDistributive(X,A,M)"
  and A2: "equiv(X,r)" 
  and A3: "Congruent2(r,A)" "Congruent2(r,M)"
  shows "IsDistributive(X//r,ProjFun2(X,r,A),ProjFun2(X,r,M))"
proof-
 let ?Ap = "ProjFun2(X,r,A)" 
 let ?Mp = "ProjFun2(X,r,M)"
 from A1 A2 A3 have
   "aX//r.bX//r.cX//r.
   ?Mp`a,?Ap`b,c = ?Ap`?Mp`a,b,?Mp`a,c  
   ?Mp`?Ap`b,c,a = ?Ap`?Mp`b,a,?Mp`c,a"
   using EquivClass_2_L3 by simp
 then show ?thesis using IsDistributive_def by simp
qed

subsectionSaturated sets

textIn this section we consider sets that are saturated 
  with respect to an equivalence
  relation. A set $A$ is saturated with respect to a relation 
  $r$ if $A=r^{-1}(r(A))$. 
  For equivalence relations saturated sets are unions of equivalence classes.
  This makes them useful as a tool to define subsets of 
  the quotient space using properties
  of representants. Namely, we often define a set 
  $B\subseteq X/r$ by saying that
  $[x]_r \in B$ iff $x\in A$. 
  If $A$ is a saturated set, this definition is consistent in 
  the sense that it does not depend on the choice of $x$ to 
  represent $[x]_r$.

textThe following defines the notion of a saturated set. 
  Recall that in Isabelle 
  r-``(A)› is the inverse image of $A$ with respect to relation $r$. 
  This definition is not specific to equivalence relations.

definition
  "IsSaturated(r,A)  A = r-``(r``(A))"

textFor equivalence relations a set is saturated iff it is an image 
  of itself.

lemma EquivClass_3_L1: assumes A1: "equiv(X,r)"
  shows "IsSaturated(r,A)  A = r``(A)"
proof
  assume "IsSaturated(r,A)"
  then have "A = (converse(r) O r)``(A)"
    using IsSaturated_def vimage_def image_comp
    by simp
  also from A1 have " = r``(A)"
    using equiv_comp_eq by simp
  finally show "A = r``(A)" by simp
next assume "A = r``(A)"
  with A1 have "A = (converse(r) O r)``(A)"
    using equiv_comp_eq by simp
  also have " =  r-``(r``(A))"
    using vimage_def image_comp by simp
  finally have "A =  r-``(r``(A))" by simp
  then show "IsSaturated(r,A)" using IsSaturated_def
    by simp
qed
  
textFor equivalence relations sets are contained in their images. 

lemma EquivClass_3_L2: assumes A1: "equiv(X,r)" and A2: "AX"
  shows "A  r``(A)"
proof
  fix a assume "aA"
  with A1 A2 have "a  r``{a}"
    using equiv_class_self by auto
  with aA show "a  r``(A)" by auto
qed
  
textThe next lemma shows that if "$\sim$" 
  is an equivalence relation and a set 
  $A$ is such that $a\in A$ and $a\sim b$ implies $b\in A$, then
  $A$ is saturated with respect to the relation.

lemma EquivClass_3_L3: assumes A1: "equiv(X,r)"
  and A2: "r  X×X" and A3: "AX"
  and A4: "xA. yX. x,y  r  yA"
  shows "IsSaturated(r,A)"
proof -
  from A2 A4 have "r``(A)  A"
    using image_iff by blast
  moreover from A1 A3 have "A  r``(A)"
    using EquivClass_3_L2 by simp
  ultimately have "A = r``(A)" by auto
  with A1 show "IsSaturated(r,A)" using EquivClass_3_L1
    by simp
qed

textIf $A\subseteq X$ and $A$ is saturated and $x\sim y$, then $x\in A$ iff
  $y\in A$. Here we show only one direction.

lemma EquivClass_3_L4: assumes A1: "equiv(X,r)"
  and A2: "IsSaturated(r,A)" and A3: "AX"
  and A4: "x,y  r" 
  and A5: "xX"  "yA"
  shows "xA"
proof -
  from A1 A5 have "x  r``{x}"
    using equiv_class_self by simp
  with A1 A3 A4 A5 have "x  r``(A)"
    using equiv_class_eq equiv_class_self
    by auto
  with A1 A2 show "xA"
    using EquivClass_3_L1 by simp
qed

textIf $A\subseteq X$ and $A$ is saturated and $x\sim y$, then $x\in A$ iff
  $y\in A$.

lemma EquivClass_3_L5: assumes A1: "equiv(X,r)"
  and A2: "IsSaturated(r,A)" and A3: "AX"
  and A4: "xX"  "yX"
  and A5: "x,y  r"
  shows "xA  yA"
proof
  assume "yA" 
  with assms show "xA" using EquivClass_3_L4
    by simp
next assume "xA"
  from A1 A5 have "y,x  r"
    using equiv_is_sym by blast
  with A1 A2 A3 A4 xA show "yA"
    using EquivClass_3_L4 by simp
qed
  
textIf $A$ is saturated then $x\in A$ iff its class is in the projection 
  of $A$.

lemma EquivClass_3_L6: assumes A1: "equiv(X,r)"
  and A2: "IsSaturated(r,A)" and A3: "AX" and A4: "xX"
  and A5: "B = {r``{x}. xA}"
  shows "xA  r``{x}  B"
proof
  assume "xA"
  with A5 show "r``{x}  B" by auto
next assume "r``{x}  B"
  with A5 obtain y where "y  A" and "r``{x} = r``{y}"
    by auto
  with A1 A3 have "x,y  r"
    using eq_equiv_class by auto
  with A1 A2 A3 A4  y  A show "xA"
    using EquivClass_3_L4 by simp
qed

textA technical lemma involving a projection of a saturated set and a 
  logical epression with exclusive or. Note that we don't really care 
  what Xor› is here, this is true for any predicate.

lemma EquivClass_3_L7: assumes "equiv(X,r)"
  and "IsSaturated(r,A)" and "AX"
  and "xX"  "yX"
  and "B = {r``{x}. xA}"
  and "(xA) Xor (yA)"
  shows "(r``{x}  B)  Xor (r``{y}  B)"
  using assms EquivClass_3_L6 by simp
    
end