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