src/HOL/Integ/Equiv.ML
changeset 3718 d78cf498a88c
parent 3457 a8ab7c64817c
child 4089 96fba19bcbe2
equal deleted inserted replaced
3717:e28553315355 3718:d78cf498a88c
    16 
    16 
    17 (** first half: equiv A r ==> r^-1 O r = r **)
    17 (** first half: equiv A r ==> r^-1 O r = r **)
    18 
    18 
    19 goalw Equiv.thy [trans_def,sym_def,inverse_def]
    19 goalw Equiv.thy [trans_def,sym_def,inverse_def]
    20     "!!r. [| sym(r); trans(r) |] ==> r^-1 O r <= r";
    20     "!!r. [| sym(r); trans(r) |] ==> r^-1 O r <= r";
    21 by (fast_tac (!claset addSEs [inverseD]) 1);
    21 by (blast_tac (!claset addSEs [inverseD]) 1);
    22 qed "sym_trans_comp_subset";
    22 qed "sym_trans_comp_subset";
    23 
    23 
    24 goalw Equiv.thy [refl_def]
    24 goalw Equiv.thy [refl_def]
    25     "!!A r. refl A r ==> r <= r^-1 O r";
    25     "!!A r. refl A r ==> r <= r^-1 O r";
    26 by (fast_tac (!claset addIs [compI]) 1);
    26 by (Blast_tac 1);
    27 qed "refl_comp_subset";
    27 qed "refl_comp_subset";
    28 
    28 
    29 goalw Equiv.thy [equiv_def]
    29 goalw Equiv.thy [equiv_def]
    30     "!!A r. equiv A r ==> r^-1 O r = r";
    30     "!!A r. equiv A r ==> r^-1 O r = r";
       
    31 by (Clarify_tac 1);
    31 by (rtac equalityI 1);
    32 by (rtac equalityI 1);
    32 by (REPEAT (ares_tac [sym_trans_comp_subset, refl_comp_subset] 1
    33 by (REPEAT (ares_tac [sym_trans_comp_subset, refl_comp_subset] 1));
    33      ORELSE etac conjE 1));
       
    34 qed "equiv_comp_eq";
    34 qed "equiv_comp_eq";
    35 
    35 
    36 (*second half*)
    36 (*second half*)
    37 goalw Equiv.thy [equiv_def,refl_def,sym_def,trans_def]
    37 goalw Equiv.thy [equiv_def,refl_def,sym_def,trans_def]
    38     "!!A r. [| r^-1 O r = r;  Domain(r) = A |] ==> equiv A r";
    38     "!!A r. [| r^-1 O r = r;  Domain(r) = A |] ==> equiv A r";
    39 by (etac equalityE 1);
    39 by (etac equalityE 1);
    40 by (subgoal_tac "ALL x y. (x,y) : r --> (y,x) : r" 1);
    40 by (subgoal_tac "ALL x y. (x,y) : r --> (y,x) : r" 1);
    41 by (Step_tac 1);
    41 by (ALLGOALS Fast_tac);
    42 by (fast_tac (!claset addIs [compI]) 3);
       
    43 by (ALLGOALS (fast_tac (!claset addIs [compI])));
       
    44 qed "comp_equivI";
    42 qed "comp_equivI";
    45 
    43 
    46 (** Equivalence classes **)
    44 (** Equivalence classes **)
    47 
    45 
    48 (*Lemma for the next result*)
    46 (*Lemma for the next result*)
    49 goalw Equiv.thy [equiv_def,trans_def,sym_def]
    47 goalw Equiv.thy [equiv_def,trans_def,sym_def]
    50     "!!A r. [| equiv A r;  (a,b): r |] ==> r^^{a} <= r^^{b}";
    48     "!!A r. [| equiv A r;  (a,b): r |] ==> r^^{a} <= r^^{b}";
    51 by (Step_tac 1);
    49 by (Blast_tac 1);
    52 by (rtac ImageI 1);
       
    53 by (Fast_tac 2);
       
    54 by (Fast_tac 1);
       
    55 qed "equiv_class_subset";
    50 qed "equiv_class_subset";
    56 
    51 
    57 goal Equiv.thy "!!A r. [| equiv A r;  (a,b): r |] ==> r^^{a} = r^^{b}";
    52 goal Equiv.thy "!!A r. [| equiv A r;  (a,b): r |] ==> r^^{a} = r^^{b}";
    58 by (REPEAT (ares_tac [equalityI, equiv_class_subset] 1));
    53 by (REPEAT (ares_tac [equalityI, equiv_class_subset] 1));
    59 by (rewrite_goals_tac [equiv_def,sym_def]);
    54 by (rewrite_goals_tac [equiv_def,sym_def]);
    60 by (Fast_tac 1);
    55 by (Blast_tac 1);
    61 qed "equiv_class_eq";
    56 qed "equiv_class_eq";
    62 
    57 
    63 goalw Equiv.thy [equiv_def,refl_def]
    58 goalw Equiv.thy [equiv_def,refl_def]
    64     "!!A r. [| equiv A r;  a: A |] ==> a: r^^{a}";
    59     "!!A r. [| equiv A r;  a: A |] ==> a: r^^{a}";
    65 by (Fast_tac 1);
    60 by (Blast_tac 1);
    66 qed "equiv_class_self";
    61 qed "equiv_class_self";
    67 
    62 
    68 (*Lemma for the next result*)
    63 (*Lemma for the next result*)
    69 goalw Equiv.thy [equiv_def,refl_def]
    64 goalw Equiv.thy [equiv_def,refl_def]
    70     "!!A r. [| equiv A r;  r^^{b} <= r^^{a};  b: A |] ==> (a,b): r";
    65     "!!A r. [| equiv A r;  r^^{b} <= r^^{a};  b: A |] ==> (a,b): r";
    71 by (Fast_tac 1);
    66 by (Blast_tac 1);
    72 qed "subset_equiv_class";
    67 qed "subset_equiv_class";
    73 
    68 
    74 goal Equiv.thy
    69 goal Equiv.thy
    75     "!!A r. [| r^^{a} = r^^{b};  equiv A r;  b: A |] ==> (a,b): r";
    70     "!!A r. [| r^^{a} = r^^{b};  equiv A r;  b: A |] ==> (a,b): r";
    76 by (REPEAT (ares_tac [equalityD2, subset_equiv_class] 1));
    71 by (REPEAT (ares_tac [equalityD2, subset_equiv_class] 1));
    77 qed "eq_equiv_class";
    72 qed "eq_equiv_class";
    78 
    73 
    79 (*thus r^^{a} = r^^{b} as well*)
    74 (*thus r^^{a} = r^^{b} as well*)
    80 goalw Equiv.thy [equiv_def,trans_def,sym_def]
    75 goalw Equiv.thy [equiv_def,trans_def,sym_def]
    81     "!!A r. [| equiv A r;  x: (r^^{a} Int r^^{b}) |] ==> (a,b): r";
    76     "!!A r. [| equiv A r;  x: (r^^{a} Int r^^{b}) |] ==> (a,b): r";
    82 by (Fast_tac 1);
    77 by (Blast_tac 1);
    83 qed "equiv_class_nondisjoint";
    78 qed "equiv_class_nondisjoint";
    84 
    79 
    85 val [major] = goalw Equiv.thy [equiv_def,refl_def]
    80 val [major] = goalw Equiv.thy [equiv_def,refl_def]
    86     "equiv A r ==> r <= A Times A";
    81     "equiv A r ==> r <= A Times A";
    87 by (rtac (major RS conjunct1 RS conjunct1) 1);
    82 by (rtac (major RS conjunct1 RS conjunct1) 1);
    88 qed "equiv_type";
    83 qed "equiv_type";
    89 
    84 
    90 goal Equiv.thy
    85 goal Equiv.thy
    91     "!!A r. equiv A r ==> ((x,y): r) = (r^^{x} = r^^{y} & x:A & y:A)";
    86     "!!A r. equiv A r ==> ((x,y): r) = (r^^{x} = r^^{y} & x:A & y:A)";
    92 by (Step_tac 1);
    87 by (blast_tac (!claset addSIs [equiv_class_eq]
    93 by ((rtac equiv_class_eq 1) THEN (assume_tac 1) THEN (assume_tac 1));
    88 	               addDs [eq_equiv_class, equiv_type]) 1);
    94 by ((rtac eq_equiv_class 3) THEN 
       
    95     (assume_tac 4) THEN (assume_tac 4) THEN (assume_tac 3));
       
    96 by ((dtac equiv_type 1) THEN (dtac rev_subsetD 1) THEN
       
    97     (assume_tac 1) THEN (dtac SigmaD1 1) THEN (assume_tac 1));
       
    98 by ((dtac equiv_type 1) THEN (dtac rev_subsetD 1) THEN
       
    99     (assume_tac 1) THEN (dtac SigmaD2 1) THEN (assume_tac 1));
       
   100 qed "equiv_class_eq_iff";
    89 qed "equiv_class_eq_iff";
   101 
    90 
   102 goal Equiv.thy
    91 goal Equiv.thy
   103     "!!A r. [| equiv A r;  x: A;  y: A |] ==> (r^^{x} = r^^{y}) = ((x,y): r)";
    92     "!!A r. [| equiv A r;  x: A;  y: A |] ==> (r^^{x} = r^^{y}) = ((x,y): r)";
   104 by (Step_tac 1);
    93 by (blast_tac (!claset addSIs [equiv_class_eq]
   105 by ((rtac eq_equiv_class 1) THEN 
    94 	               addDs [eq_equiv_class, equiv_type]) 1);
   106     (assume_tac 1) THEN (assume_tac 1) THEN (assume_tac 1));
       
   107 by ((rtac equiv_class_eq 1) THEN 
       
   108     (assume_tac 1) THEN (assume_tac 1));
       
   109 qed "eq_equiv_class_iff";
    95 qed "eq_equiv_class_iff";
   110 
    96 
   111 (*** Quotients ***)
    97 (*** Quotients ***)
   112 
    98 
   113 (** Introduction/elimination rules -- needed? **)
    99 (** Introduction/elimination rules -- needed? **)
   114 
   100 
   115 goalw Equiv.thy [quotient_def] "!!A. x:A ==> r^^{x}: A/r";
   101 goalw Equiv.thy [quotient_def] "!!A. x:A ==> r^^{x}: A/r";
   116 by (Fast_tac 1);
   102 by (Blast_tac 1);
   117 qed "quotientI";
   103 qed "quotientI";
   118 
   104 
   119 val [major,minor] = goalw Equiv.thy [quotient_def]
   105 val [major,minor] = goalw Equiv.thy [quotient_def]
   120     "[| X:(A/r);  !!x. [| X = r^^{x};  x:A |] ==> P |]  \
   106     "[| X:(A/r);  !!x. [| X = r^^{x};  x:A |] ==> P |]  \
   121 \    ==> P";
   107 \    ==> P";
   122 by (resolve_tac [major RS UN_E] 1);
   108 by (resolve_tac [major RS UN_E] 1);
   123 by (rtac minor 1);
   109 by (rtac minor 1);
   124 by (assume_tac 2);
   110 by (assume_tac 2);
   125 by (Fast_tac 1);
   111 by (Fast_tac 1);   (*Blast_tac FAILS to prove it*)
   126 qed "quotientE";
   112 qed "quotientE";
   127 
   113 
   128 goalw Equiv.thy [equiv_def,refl_def,quotient_def]
   114 goalw Equiv.thy [equiv_def,refl_def,quotient_def]
   129     "!!A r. equiv A r ==> Union(A/r) = A";
   115     "!!A r. equiv A r ==> Union(A/r) = A";
   130 by (blast_tac (!claset addSIs [equalityI]) 1);
   116 by (blast_tac (!claset addSIs [equalityI]) 1);
   155 (*Conversion rule*)
   141 (*Conversion rule*)
   156 goal Equiv.thy "!!A r. [| equiv A r;  congruent r b;  a: A |] \
   142 goal Equiv.thy "!!A r. [| equiv A r;  congruent r b;  a: A |] \
   157 \                      ==> (UN x:r^^{a}. b(x)) = b(a)";
   143 \                      ==> (UN x:r^^{a}. b(x)) = b(a)";
   158 by (rtac (equiv_class_self RS UN_singleton) 1 THEN REPEAT (assume_tac 1));
   144 by (rtac (equiv_class_self RS UN_singleton) 1 THEN REPEAT (assume_tac 1));
   159 by (rewrite_goals_tac [equiv_def,congruent_def,sym_def]);
   145 by (rewrite_goals_tac [equiv_def,congruent_def,sym_def]);
   160 by (Fast_tac 1);
   146 by (Blast_tac 1);
   161 qed "UN_equiv_class";
   147 qed "UN_equiv_class";
   162 
   148 
   163 (*type checking of  UN x:r``{a}. b(x) *)
   149 (*type checking of  UN x:r``{a}. b(x) *)
   164 val prems = goalw Equiv.thy [quotient_def]
   150 val prems = goalw Equiv.thy [quotient_def]
   165     "[| equiv A r;  congruent r b;  X: A/r;     \
   151     "[| equiv A r;  congruent r b;  X: A/r;     \
   166 \       !!x.  x : A ==> b(x) : B |]             \
   152 \       !!x.  x : A ==> b(x) : B |]             \
   167 \    ==> (UN x:X. b(x)) : B";
   153 \    ==> (UN x:X. b(x)) : B";
   168 by (cut_facts_tac prems 1);
   154 by (cut_facts_tac prems 1);
   169 by (Step_tac 1);
   155 by (Clarify_tac 1);
   170 by (stac UN_equiv_class 1);
   156 by (stac UN_equiv_class 1);
   171 by (REPEAT (ares_tac prems 1));
   157 by (REPEAT (ares_tac prems 1));
   172 qed "UN_equiv_class_type";
   158 qed "UN_equiv_class_type";
   173 
   159 
   174 (*Sufficient conditions for injectiveness.  Could weaken premises!
   160 (*Sufficient conditions for injectiveness.  Could weaken premises!
   178     "[| equiv A r;   congruent r b;  \
   164     "[| equiv A r;   congruent r b;  \
   179 \       (UN x:X. b(x))=(UN y:Y. b(y));  X: A/r;  Y: A/r;        \
   165 \       (UN x:X. b(x))=(UN y:Y. b(y));  X: A/r;  Y: A/r;        \
   180 \       !!x y. [| x:A; y:A; b(x)=b(y) |] ==> (x,y):r |]         \
   166 \       !!x y. [| x:A; y:A; b(x)=b(y) |] ==> (x,y):r |]         \
   181 \    ==> X=Y";
   167 \    ==> X=Y";
   182 by (cut_facts_tac prems 1);
   168 by (cut_facts_tac prems 1);
   183 by (Step_tac 1);
   169 by (Clarify_tac 1);
   184 by (rtac equiv_class_eq 1);
   170 by (rtac equiv_class_eq 1);
   185 by (REPEAT (ares_tac prems 1));
   171 by (REPEAT (ares_tac prems 1));
   186 by (etac box_equals 1);
   172 by (etac box_equals 1);
   187 by (REPEAT (ares_tac [UN_equiv_class] 1));
   173 by (REPEAT (ares_tac [UN_equiv_class] 1));
   188 qed "UN_equiv_class_inject";
   174 qed "UN_equiv_class_inject";
   191 (**** Defining binary operations upon equivalence classes ****)
   177 (**** Defining binary operations upon equivalence classes ****)
   192 
   178 
   193 
   179 
   194 goalw Equiv.thy [congruent_def,congruent2_def,equiv_def,refl_def]
   180 goalw Equiv.thy [congruent_def,congruent2_def,equiv_def,refl_def]
   195     "!!A r. [| equiv A r;  congruent2 r b;  a: A |] ==> congruent r (b a)";
   181     "!!A r. [| equiv A r;  congruent2 r b;  a: A |] ==> congruent r (b a)";
   196 by (Fast_tac 1);
   182 by (Blast_tac 1);
   197 qed "congruent2_implies_congruent";
   183 qed "congruent2_implies_congruent";
   198 
   184 
   199 goalw Equiv.thy [congruent_def]
   185 goalw Equiv.thy [congruent_def]
   200     "!!A r. [| equiv A r;  congruent2 r b;  a: A |] ==> \
   186     "!!A r. [| equiv A r;  congruent2 r b;  a: A |] ==> \
   201 \    congruent r (%x1. UN x2:r^^{a}. b x1 x2)";
   187 \    congruent r (%x1. UN x2:r^^{a}. b x1 x2)";
   202 by (Step_tac 1);
   188 by (Clarify_tac 1);
   203 by (rtac (equiv_type RS subsetD RS SigmaE2) 1 THEN REPEAT (assume_tac 1));
   189 by (rtac (equiv_type RS subsetD RS SigmaE2) 1 THEN REPEAT (assume_tac 1));
   204 by (asm_simp_tac (!simpset addsimps [UN_equiv_class,
   190 by (asm_simp_tac (!simpset addsimps [UN_equiv_class,
   205                                      congruent2_implies_congruent]) 1);
   191                                      congruent2_implies_congruent]) 1);
   206 by (rewrite_goals_tac [congruent2_def,equiv_def,refl_def]);
   192 by (rewrite_goals_tac [congruent2_def,equiv_def,refl_def]);
   207 by (Fast_tac 1);
   193 by (Blast_tac 1);
   208 qed "congruent2_implies_congruent_UN";
   194 qed "congruent2_implies_congruent_UN";
   209 
   195 
   210 goal Equiv.thy
   196 goal Equiv.thy
   211     "!!A r. [| equiv A r;  congruent2 r b;  a1: A;  a2: A |]  \
   197     "!!A r. [| equiv A r;  congruent2 r b;  a1: A;  a2: A |]  \
   212 \    ==> (UN x1:r^^{a1}. UN x2:r^^{a2}. b x1 x2) = b a1 a2";
   198 \    ==> (UN x1:r^^{a1}. UN x2:r^^{a2}. b x1 x2) = b a1 a2";
   220     "[| equiv A r;  congruent2 r b;  \
   206     "[| equiv A r;  congruent2 r b;  \
   221 \       X1: A/r;  X2: A/r;      \
   207 \       X1: A/r;  X2: A/r;      \
   222 \       !!x1 x2.  [| x1: A; x2: A |] ==> b x1 x2 : B |]    \
   208 \       !!x1 x2.  [| x1: A; x2: A |] ==> b x1 x2 : B |]    \
   223 \    ==> (UN x1:X1. UN x2:X2. b x1 x2) : B";
   209 \    ==> (UN x1:X1. UN x2:X2. b x1 x2) : B";
   224 by (cut_facts_tac prems 1);
   210 by (cut_facts_tac prems 1);
   225 by (Step_tac 1);
   211 by (Clarify_tac 1);
   226 by (REPEAT (ares_tac (prems@[UN_equiv_class_type,
   212 by (REPEAT (ares_tac (prems@[UN_equiv_class_type,
   227                              congruent2_implies_congruent_UN,
   213                              congruent2_implies_congruent_UN,
   228                              congruent2_implies_congruent, quotientI]) 1));
   214                              congruent2_implies_congruent, quotientI]) 1));
   229 qed "UN_equiv_class_type2";
   215 qed "UN_equiv_class_type2";
   230 
   216 
   235     "[| equiv A r;      \
   221     "[| equiv A r;      \
   236 \       !! y z w. [| w: A;  (y,z) : r |] ==> b y w = b z w;      \
   222 \       !! y z w. [| w: A;  (y,z) : r |] ==> b y w = b z w;      \
   237 \       !! y z w. [| w: A;  (y,z) : r |] ==> b w y = b w z       \
   223 \       !! y z w. [| w: A;  (y,z) : r |] ==> b w y = b w z       \
   238 \    |] ==> congruent2 r b";
   224 \    |] ==> congruent2 r b";
   239 by (cut_facts_tac prems 1);
   225 by (cut_facts_tac prems 1);
   240 by (Step_tac 1);
   226 by (Clarify_tac 1);
   241 by (rtac trans 1);
   227 by (blast_tac (!claset addIs (trans::prems)) 1);
   242 by (REPEAT (ares_tac prems 1
       
   243      ORELSE etac (subsetD RS SigmaE2) 1 THEN assume_tac 2 THEN assume_tac 1));
       
   244 qed "congruent2I";
   228 qed "congruent2I";
   245 
   229 
   246 val [equivA,commute,congt] = goal Equiv.thy
   230 val [equivA,commute,congt] = goal Equiv.thy
   247     "[| equiv A r;      \
   231     "[| equiv A r;      \
   248 \       !! y z. [| y: A;  z: A |] ==> b y z = b z y;        \
   232 \       !! y z. [| y: A;  z: A |] ==> b y z = b z y;        \