src/ZF/Integ/EquivClass.ML
changeset 5528 4896b4e4077b
child 6153 bff90585cce5
equal deleted inserted replaced
5527:38928c4a8eb2 5528:4896b4e4077b
       
     1 (*  Title:      ZF/EquivClass.ML
       
     2     ID:         $Id$
       
     3     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
       
     4     Copyright   1994  University of Cambridge
       
     5 
       
     6 Equivalence relations in Zermelo-Fraenkel Set Theory 
       
     7 *)
       
     8 
       
     9 val RSLIST = curry (op MRS);
       
    10 
       
    11 open EquivClass;
       
    12 
       
    13 (*** Suppes, Theorem 70: r is an equiv relation iff converse(r) O r = r ***)
       
    14 
       
    15 (** first half: equiv(A,r) ==> converse(r) O r = r **)
       
    16 
       
    17 Goalw [trans_def,sym_def]
       
    18     "[| sym(r); trans(r) |] ==> converse(r) O r <= r";
       
    19 by (Blast_tac 1);
       
    20 qed "sym_trans_comp_subset";
       
    21 
       
    22 Goalw [refl_def]
       
    23     "[| refl(A,r); r <= A*A |] ==> r <= converse(r) O r";
       
    24 by (Blast_tac 1);
       
    25 qed "refl_comp_subset";
       
    26 
       
    27 Goalw [equiv_def]
       
    28     "equiv(A,r) ==> converse(r) O r = r";
       
    29 by (blast_tac (subset_cs addSIs [sym_trans_comp_subset, refl_comp_subset]) 1);
       
    30 qed "equiv_comp_eq";
       
    31 
       
    32 (*second half*)
       
    33 Goalw [equiv_def,refl_def,sym_def,trans_def]
       
    34     "[| converse(r) O r = r;  domain(r) = A |] ==> equiv(A,r)";
       
    35 by (etac equalityE 1);
       
    36 by (subgoal_tac "ALL x y. <x,y> : r --> <y,x> : r" 1);
       
    37 by (ALLGOALS Fast_tac);
       
    38 qed "comp_equivI";
       
    39 
       
    40 (** Equivalence classes **)
       
    41 
       
    42 (*Lemma for the next result*)
       
    43 Goalw [trans_def,sym_def]
       
    44     "[| sym(r);  trans(r);  <a,b>: r |] ==> r``{a} <= r``{b}";
       
    45 by (Blast_tac 1);
       
    46 qed "equiv_class_subset";
       
    47 
       
    48 Goalw [equiv_def]
       
    49     "[| equiv(A,r);  <a,b>: r |] ==> r``{a} = r``{b}";
       
    50 by (safe_tac (subset_cs addSIs [equalityI, equiv_class_subset]));
       
    51 by (rewtac sym_def);
       
    52 by (Blast_tac 1);
       
    53 qed "equiv_class_eq";
       
    54 
       
    55 Goalw [equiv_def,refl_def]
       
    56     "[| equiv(A,r);  a: A |] ==> a: r``{a}";
       
    57 by (Blast_tac 1);
       
    58 qed "equiv_class_self";
       
    59 
       
    60 (*Lemma for the next result*)
       
    61 Goalw [equiv_def,refl_def]
       
    62     "[| equiv(A,r);  r``{b} <= r``{a};  b: A |] ==> <a,b>: r";
       
    63 by (Blast_tac 1);
       
    64 qed "subset_equiv_class";
       
    65 
       
    66 Goal "[| r``{a} = r``{b};  equiv(A,r);  b: A |] ==> <a,b>: r";
       
    67 by (REPEAT (ares_tac[equalityD2, subset_equiv_class] 1));
       
    68 qed "eq_equiv_class";
       
    69 
       
    70 (*thus r``{a} = r``{b} as well*)
       
    71 Goalw [equiv_def,trans_def,sym_def]
       
    72     "[| equiv(A,r);  x: (r``{a} Int r``{b}) |] ==> <a,b>: r";
       
    73 by (Blast_tac 1);
       
    74 qed "equiv_class_nondisjoint";
       
    75 
       
    76 Goalw [equiv_def] "equiv(A,r) ==> r <= A*A";
       
    77 by (safe_tac subset_cs);
       
    78 qed "equiv_type";
       
    79 
       
    80 Goal "equiv(A,r) ==> <x,y>: r <-> r``{x} = r``{y} & x:A & y:A";
       
    81 by (blast_tac (claset() addIs [eq_equiv_class, equiv_class_eq]
       
    82                       addDs [equiv_type]) 1);
       
    83 qed "equiv_class_eq_iff";
       
    84 
       
    85 Goal "[| equiv(A,r);  x: A;  y: A |] ==> r``{x} = r``{y} <-> <x,y>: r";
       
    86 by (blast_tac (claset() addIs [eq_equiv_class, equiv_class_eq]
       
    87                       addDs [equiv_type]) 1);
       
    88 qed "eq_equiv_class_iff";
       
    89 
       
    90 (*** Quotients ***)
       
    91 
       
    92 (** Introduction/elimination rules -- needed? **)
       
    93 
       
    94 Goalw [quotient_def] "x:A ==> r``{x}: A/r";
       
    95 by (etac RepFunI 1);
       
    96 qed "quotientI";
       
    97 
       
    98 val major::prems = Goalw [quotient_def]
       
    99     "[| X: A/r;  !!x. [| X = r``{x};  x:A |] ==> P |]   \
       
   100 \    ==> P";
       
   101 by (rtac (major RS RepFunE) 1);
       
   102 by (eresolve_tac prems 1);
       
   103 by (assume_tac 1);
       
   104 qed "quotientE";
       
   105 
       
   106 Goalw [equiv_def,refl_def,quotient_def]
       
   107     "equiv(A,r) ==> Union(A/r) = A";
       
   108 by (Blast_tac 1);
       
   109 qed "Union_quotient";
       
   110 
       
   111 Goalw [quotient_def]
       
   112     "[| equiv(A,r);  X: A/r;  Y: A/r |] ==> X=Y | (X Int Y <= 0)";
       
   113 by (safe_tac (claset() addSIs [equiv_class_eq]));
       
   114 by (assume_tac 1);
       
   115 by (rewrite_goals_tac [equiv_def,trans_def,sym_def]);
       
   116 by (Blast_tac 1);
       
   117 qed "quotient_disj";
       
   118 
       
   119 (**** Defining unary operations upon equivalence classes ****)
       
   120 
       
   121 (** These proofs really require as local premises
       
   122      equiv(A,r);  congruent(r,b)
       
   123 **)
       
   124 
       
   125 (*Conversion rule*)
       
   126 val prems as [equivA,bcong,_] = goal EquivClass.thy
       
   127     "[| equiv(A,r);  congruent(r,b);  a: A |] ==> (UN x:r``{a}. b(x)) = b(a)";
       
   128 by (cut_facts_tac prems 1);
       
   129 by (rtac ([refl RS UN_cong, UN_constant] MRS trans) 1);
       
   130 by (etac equiv_class_self 2);
       
   131 by (assume_tac 2);
       
   132 by (rewrite_goals_tac [equiv_def,sym_def,congruent_def]);
       
   133 by (Blast_tac 1);
       
   134 qed "UN_equiv_class";
       
   135 
       
   136 (*type checking of  UN x:r``{a}. b(x) *)
       
   137 val prems = Goalw [quotient_def]
       
   138     "[| equiv(A,r);  congruent(r,b);  X: A/r;   \
       
   139 \       !!x.  x : A ==> b(x) : B |]     \
       
   140 \    ==> (UN x:X. b(x)) : B";
       
   141 by (cut_facts_tac prems 1);
       
   142 by Safe_tac;
       
   143 by (asm_simp_tac (simpset() addsimps UN_equiv_class::prems) 1);
       
   144 qed "UN_equiv_class_type";
       
   145 
       
   146 (*Sufficient conditions for injectiveness.  Could weaken premises!
       
   147   major premise could be an inclusion; bcong could be !!y. y:A ==> b(y):B
       
   148 *)
       
   149 val prems = Goalw [quotient_def]
       
   150     "[| equiv(A,r);   congruent(r,b);  \
       
   151 \       (UN x:X. b(x))=(UN y:Y. b(y));  X: A/r;  Y: A/r;  \
       
   152 \       !!x y. [| x:A; y:A; b(x)=b(y) |] ==> <x,y>:r |]         \
       
   153 \    ==> X=Y";
       
   154 by (cut_facts_tac prems 1);
       
   155 by Safe_tac;
       
   156 by (rtac equiv_class_eq 1);
       
   157 by (REPEAT (ares_tac prems 1));
       
   158 by (etac box_equals 1);
       
   159 by (REPEAT (ares_tac [UN_equiv_class] 1));
       
   160 qed "UN_equiv_class_inject";
       
   161 
       
   162 
       
   163 (**** Defining binary operations upon equivalence classes ****)
       
   164 
       
   165 
       
   166 Goalw [congruent_def,congruent2_def,equiv_def,refl_def]
       
   167     "[| equiv(A,r);  congruent2(r,b);  a: A |] ==> congruent(r,b(a))";
       
   168 by (Blast_tac 1);
       
   169 qed "congruent2_implies_congruent";
       
   170 
       
   171 val equivA::prems = goalw EquivClass.thy [congruent_def]
       
   172     "[| equiv(A,r);  congruent2(r,b);  a: A |] ==> \
       
   173 \    congruent(r, %x1. UN x2:r``{a}. b(x1,x2))";
       
   174 by (cut_facts_tac (equivA::prems) 1);
       
   175 by Safe_tac;
       
   176 by (rtac (equivA RS equiv_type RS subsetD RS SigmaE2) 1);
       
   177 by (assume_tac 1);
       
   178 by (asm_simp_tac (simpset() addsimps [equivA RS UN_equiv_class,
       
   179                                      congruent2_implies_congruent]) 1);
       
   180 by (rewrite_goals_tac [congruent2_def,equiv_def,refl_def]);
       
   181 by (Blast_tac 1);
       
   182 qed "congruent2_implies_congruent_UN";
       
   183 
       
   184 val prems as equivA::_ = goal EquivClass.thy
       
   185     "[| equiv(A,r);  congruent2(r,b);  a1: A;  a2: A |]  \
       
   186 \    ==> (UN x1:r``{a1}. UN x2:r``{a2}. b(x1,x2)) = b(a1,a2)";
       
   187 by (cut_facts_tac prems 1);
       
   188 by (asm_simp_tac (simpset() addsimps [equivA RS UN_equiv_class,
       
   189                                      congruent2_implies_congruent,
       
   190                                      congruent2_implies_congruent_UN]) 1);
       
   191 qed "UN_equiv_class2";
       
   192 
       
   193 (*type checking*)
       
   194 val prems = Goalw [quotient_def]
       
   195     "[| equiv(A,r);  congruent2(r,b);                   \
       
   196 \       X1: A/r;  X2: A/r;                              \
       
   197 \       !!x1 x2.  [| x1: A; x2: A |] ==> b(x1,x2) : B   \
       
   198 \    |] ==> (UN x1:X1. UN x2:X2. b(x1,x2)) : B";
       
   199 by (cut_facts_tac prems 1);
       
   200 by Safe_tac;
       
   201 by (REPEAT (ares_tac (prems@[UN_equiv_class_type,
       
   202                              congruent2_implies_congruent_UN,
       
   203                              congruent2_implies_congruent, quotientI]) 1));
       
   204 qed "UN_equiv_class_type2";
       
   205 
       
   206 
       
   207 (*Suggested by John Harrison -- the two subproofs may be MUCH simpler
       
   208   than the direct proof*)
       
   209 val prems = Goalw [congruent2_def,equiv_def,refl_def]
       
   210     "[| equiv(A,r);     \
       
   211 \       !! y z w. [| w: A;  <y,z> : r |] ==> b(y,w) = b(z,w);      \
       
   212 \       !! y z w. [| w: A;  <y,z> : r |] ==> b(w,y) = b(w,z)       \
       
   213 \    |] ==> congruent2(r,b)";
       
   214 by (cut_facts_tac prems 1);
       
   215 by Safe_tac;
       
   216 by (rtac trans 1);
       
   217 by (REPEAT (ares_tac prems 1
       
   218      ORELSE etac (subsetD RS SigmaE2) 1 THEN assume_tac 2 THEN assume_tac 1));
       
   219 qed "congruent2I";
       
   220 
       
   221 val [equivA,commute,congt] = Goal
       
   222     "[| equiv(A,r);     \
       
   223 \       !! y z. [| y: A;  z: A |] ==> b(y,z) = b(z,y);        \
       
   224 \       !! y z w. [| w: A;  <y,z>: r |] ==> b(w,y) = b(w,z)     \
       
   225 \    |] ==> congruent2(r,b)";
       
   226 by (resolve_tac [equivA RS congruent2I] 1);
       
   227 by (rtac (commute RS trans) 1);
       
   228 by (rtac (commute RS trans RS sym) 3);
       
   229 by (rtac sym 5);
       
   230 by (REPEAT (ares_tac [congt] 1
       
   231      ORELSE etac (equivA RS equiv_type RS subsetD RS SigmaE2) 1));
       
   232 qed "congruent2_commuteI";
       
   233 
       
   234 (*Obsolete?*)
       
   235 val [equivA,ZinA,congt,commute] = Goalw [quotient_def]
       
   236     "[| equiv(A,r);  Z: A/r;  \
       
   237 \       !!w. [| w: A |] ==> congruent(r, %z. b(w,z));    \
       
   238 \       !!x y. [| x: A;  y: A |] ==> b(y,x) = b(x,y)    \
       
   239 \    |] ==> congruent(r, %w. UN z: Z. b(w,z))";
       
   240 val congt' = rewrite_rule [congruent_def] congt;
       
   241 by (cut_facts_tac [ZinA] 1);
       
   242 by (rewtac congruent_def);
       
   243 by Safe_tac;
       
   244 by (rtac (equivA RS equiv_type RS subsetD RS SigmaE2) 1);
       
   245 by (assume_tac 1);
       
   246 by (asm_simp_tac (simpset() addsimps [commute,
       
   247                                      [equivA, congt] MRS UN_equiv_class]) 1);
       
   248 by (REPEAT (ares_tac [congt' RS spec RS spec RS mp] 1));
       
   249 qed "congruent_commuteI";