src/ZF/EquivClass.ML
author paulson
Fri, 21 Aug 1998 16:14:34 +0200
changeset 5359 bd539b72d484
parent 5268 59ef39008514
permissions -rw-r--r--
Tidying
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1076
diff changeset
     1
(*  Title:      ZF/EquivClass.ML
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1076
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
535
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
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5147
diff changeset
     6
Equivalence relations in Zermelo-Fraenkel Set Theory 
535
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
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
    17
Goalw [trans_def,sym_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
    18
    "[| sym(r); trans(r) |] ==> converse(r) O r <= r";
2929
4eefc6c22d41 Changed some fast_tac to blast_tac
paulson
parents: 2496
diff changeset
    19
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 749
diff changeset
    20
qed "sym_trans_comp_subset";
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
    21
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
    22
Goalw [refl_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
    23
    "[| refl(A,r); r <= A*A |] ==> r <= converse(r) O r";
2929
4eefc6c22d41 Changed some fast_tac to blast_tac
paulson
parents: 2496
diff changeset
    24
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 749
diff changeset
    25
qed "refl_comp_subset";
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
    26
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
    27
Goalw [equiv_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
    28
    "equiv(A,r) ==> converse(r) O r = r";
2929
4eefc6c22d41 Changed some fast_tac to blast_tac
paulson
parents: 2496
diff changeset
    29
by (blast_tac (subset_cs addSIs [sym_trans_comp_subset, refl_comp_subset]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 749
diff changeset
    30
qed "equiv_comp_eq";
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
    31
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
    32
(*second half*)
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
    33
Goalw [equiv_def,refl_def,sym_def,trans_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
    34
    "[| converse(r) O r = r;  domain(r) = A |] ==> equiv(A,r)";
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
    35
by (etac equalityE 1);
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
    36
by (subgoal_tac "ALL x y. <x,y> : r --> <y,x> : r" 1);
2929
4eefc6c22d41 Changed some fast_tac to blast_tac
paulson
parents: 2496
diff changeset
    37
by (ALLGOALS Fast_tac);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 749
diff changeset
    38
qed "comp_equivI";
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
    39
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
    40
(** Equivalence classes **)
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
    41
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
    42
(*Lemma for the next result*)
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
    43
Goalw [trans_def,sym_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
    44
    "[| sym(r);  trans(r);  <a,b>: r |] ==> r``{a} <= r``{b}";
2929
4eefc6c22d41 Changed some fast_tac to blast_tac
paulson
parents: 2496
diff changeset
    45
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 749
diff changeset
    46
qed "equiv_class_subset";
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
    47
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
    48
Goalw [equiv_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
    49
    "[| equiv(A,r);  <a,b>: r |] ==> r``{a} = r``{b}";
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
    50
by (safe_tac (subset_cs addSIs [equalityI, equiv_class_subset]));
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1076
diff changeset
    51
by (rewtac sym_def);
2929
4eefc6c22d41 Changed some fast_tac to blast_tac
paulson
parents: 2496
diff changeset
    52
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 749
diff changeset
    53
qed "equiv_class_eq";
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
    54
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
    55
Goalw [equiv_def,refl_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
    56
    "[| equiv(A,r);  a: A |] ==> a: r``{a}";
2929
4eefc6c22d41 Changed some fast_tac to blast_tac
paulson
parents: 2496
diff changeset
    57
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 749
diff changeset
    58
qed "equiv_class_self";
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
    59
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
    60
(*Lemma for the next result*)
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
    61
Goalw [equiv_def,refl_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
    62
    "[| equiv(A,r);  r``{b} <= r``{a};  b: A |] ==> <a,b>: r";
2929
4eefc6c22d41 Changed some fast_tac to blast_tac
paulson
parents: 2496
diff changeset
    63
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 749
diff changeset
    64
qed "subset_equiv_class";
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
    65
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5147
diff changeset
    66
Goal "[| r``{a} = r``{b};  equiv(A,r);  b: A |] ==> <a,b>: r";
59ef39008514 even more tidying of Goal commands
paulson
parents: 5147
diff changeset
    67
by (REPEAT (ares_tac[equalityD2, subset_equiv_class] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 749
diff changeset
    68
qed "eq_equiv_class";
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
    69
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
    70
(*thus r``{a} = r``{b} as well*)
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
    71
Goalw [equiv_def,trans_def,sym_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
    72
    "[| equiv(A,r);  x: (r``{a} Int r``{b}) |] ==> <a,b>: r";
2929
4eefc6c22d41 Changed some fast_tac to blast_tac
paulson
parents: 2496
diff changeset
    73
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 749
diff changeset
    74
qed "equiv_class_nondisjoint";
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
    75
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
    76
Goalw [equiv_def] "equiv(A,r) ==> r <= A*A";
1076
68f088887f48 Modified proofs for new claset primitives. The problem is that they enforce
lcp
parents: 1013
diff changeset
    77
by (safe_tac subset_cs);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 749
diff changeset
    78
qed "equiv_type";
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
    79
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5147
diff changeset
    80
Goal "equiv(A,r) ==> <x,y>: r <-> r``{x} = r``{y} & x:A & y:A";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
    81
by (blast_tac (claset() addIs [eq_equiv_class, equiv_class_eq]
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    82
                      addDs [equiv_type]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 749
diff changeset
    83
qed "equiv_class_eq_iff";
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
    84
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5147
diff changeset
    85
Goal "[| equiv(A,r);  x: A;  y: A |] ==> r``{x} = r``{y} <-> <x,y>: r";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
    86
by (blast_tac (claset() addIs [eq_equiv_class, equiv_class_eq]
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1461
diff changeset
    87
                      addDs [equiv_type]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 749
diff changeset
    88
qed "eq_equiv_class_iff";
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
    89
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
    90
(*** Quotients ***)
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
(** Introduction/elimination rules -- needed? **)
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
    93
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5147
diff changeset
    94
Goalw [quotient_def] "x:A ==> r``{x}: A/r";
59ef39008514 even more tidying of Goal commands
paulson
parents: 5147
diff changeset
    95
by (etac RepFunI 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 749
diff changeset
    96
qed "quotientI";
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
    97
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5147
diff changeset
    98
val major::prems = Goalw [quotient_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1076
diff changeset
    99
    "[| X: A/r;  !!x. [| X = r``{x};  x:A |] ==> P |]   \
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   100
\    ==> P";
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   101
by (rtac (major RS RepFunE) 1);
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   102
by (eresolve_tac prems 1);
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   103
by (assume_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 749
diff changeset
   104
qed "quotientE";
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   105
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   106
Goalw [equiv_def,refl_def,quotient_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   107
    "equiv(A,r) ==> Union(A/r) = A";
2929
4eefc6c22d41 Changed some fast_tac to blast_tac
paulson
parents: 2496
diff changeset
   108
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 749
diff changeset
   109
qed "Union_quotient";
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   110
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   111
Goalw [quotient_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   112
    "[| equiv(A,r);  X: A/r;  Y: A/r |] ==> X=Y | (X Int Y <= 0)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   113
by (safe_tac (claset() addSIs [equiv_class_eq]));
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   114
by (assume_tac 1);
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   115
by (rewrite_goals_tac [equiv_def,trans_def,sym_def]);
2929
4eefc6c22d41 Changed some fast_tac to blast_tac
paulson
parents: 2496
diff changeset
   116
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 749
diff changeset
   117
qed "quotient_disj";
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   118
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   119
(**** Defining unary operations upon equivalence classes ****)
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   120
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   121
(** 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
   122
     equiv(A,r);  congruent(r,b)
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   123
**)
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   124
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   125
(*Conversion rule*)
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   126
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
   127
    "[| 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
   128
by (cut_facts_tac prems 1);
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   129
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
   130
by (etac equiv_class_self 2);
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   131
by (assume_tac 2);
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   132
by (rewrite_goals_tac [equiv_def,sym_def,congruent_def]);
2929
4eefc6c22d41 Changed some fast_tac to blast_tac
paulson
parents: 2496
diff changeset
   133
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 749
diff changeset
   134
qed "UN_equiv_class";
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   135
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   136
(*type checking of  UN x:r``{a}. b(x) *)
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5147
diff changeset
   137
val prems = Goalw [quotient_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1076
diff changeset
   138
    "[| equiv(A,r);  congruent(r,b);  X: A/r;   \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1076
diff changeset
   139
\       !!x.  x : A ==> b(x) : B |]     \
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   140
\    ==> (UN x:X. b(x)) : B";
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   141
by (cut_facts_tac prems 1);
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
   142
by Safe_tac;
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   143
by (asm_simp_tac (simpset() addsimps (UN_equiv_class::prems)) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 749
diff changeset
   144
qed "UN_equiv_class_type";
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   145
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   146
(*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
   147
  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
   148
*)
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5147
diff changeset
   149
val prems = Goalw [quotient_def]
535
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);  \
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   151
\       (UN x:X. b(x))=(UN y:Y. b(y));  X: A/r;  Y: A/r;  \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1076
diff changeset
   152
\       !!x y. [| x:A; y:A; b(x)=b(y) |] ==> <x,y>:r |]         \
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   153
\    ==> X=Y";
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   154
by (cut_facts_tac prems 1);
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
   155
by Safe_tac;
946
c0f4ae3fda92 Removed "localize"; instead, proofs refer to their own
lcp
parents: 760
diff changeset
   156
by (rtac equiv_class_eq 1);
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   157
by (REPEAT (ares_tac prems 1));
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   158
by (etac box_equals 1);
946
c0f4ae3fda92 Removed "localize"; instead, proofs refer to their own
lcp
parents: 760
diff changeset
   159
by (REPEAT (ares_tac [UN_equiv_class] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 749
diff changeset
   160
qed "UN_equiv_class_inject";
535
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
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   163
(**** Defining binary operations upon equivalence classes ****)
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   164
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   165
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   166
Goalw [congruent_def,congruent2_def,equiv_def,refl_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   167
    "[| equiv(A,r);  congruent2(r,b);  a: A |] ==> congruent(r,b(a))";
2929
4eefc6c22d41 Changed some fast_tac to blast_tac
paulson
parents: 2496
diff changeset
   168
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 749
diff changeset
   169
qed "congruent2_implies_congruent";
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   170
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   171
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
   172
    "[| 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
   173
\    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
   174
by (cut_facts_tac (equivA::prems) 1);
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
   175
by Safe_tac;
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   176
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
   177
by (assume_tac 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   178
by (asm_simp_tac (simpset() addsimps [equivA RS UN_equiv_class,
2493
bdeb5024353a Removal of sum_cs and eq_cs
paulson
parents: 2469
diff changeset
   179
                                     congruent2_implies_congruent]) 1);
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   180
by (rewrite_goals_tac [congruent2_def,equiv_def,refl_def]);
2929
4eefc6c22d41 Changed some fast_tac to blast_tac
paulson
parents: 2496
diff changeset
   181
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 749
diff changeset
   182
qed "congruent2_implies_congruent_UN";
535
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 prems as equivA::_ = goal EquivClass.thy
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);  a1: A;  a2: A |]  \
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   186
\    ==> (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
   187
by (cut_facts_tac prems 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   188
by (asm_simp_tac (simpset() addsimps [equivA RS UN_equiv_class,
2493
bdeb5024353a Removal of sum_cs and eq_cs
paulson
parents: 2469
diff changeset
   189
                                     congruent2_implies_congruent,
bdeb5024353a Removal of sum_cs and eq_cs
paulson
parents: 2469
diff changeset
   190
                                     congruent2_implies_congruent_UN]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 749
diff changeset
   191
qed "UN_equiv_class2";
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   192
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   193
(*type checking*)
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5147
diff changeset
   194
val prems = Goalw [quotient_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1076
diff changeset
   195
    "[| equiv(A,r);  congruent2(r,b);                   \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1076
diff changeset
   196
\       X1: A/r;  X2: A/r;                              \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1076
diff changeset
   197
\       !!x1 x2.  [| x1: A; x2: A |] ==> b(x1,x2) : B   \
749
4b605159b5a5 equiv_comp_eq: simplified proof
lcp
parents: 723
diff changeset
   198
\    |] ==> (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
   199
by (cut_facts_tac prems 1);
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
   200
by Safe_tac;
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   201
by (REPEAT (ares_tac (prems@[UN_equiv_class_type,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1076
diff changeset
   202
                             congruent2_implies_congruent_UN,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1076
diff changeset
   203
                             congruent2_implies_congruent, quotientI]) 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 749
diff changeset
   204
qed "UN_equiv_class_type2";
535
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
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   207
(*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
   208
  than the direct proof*)
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5147
diff changeset
   209
val prems = Goalw [congruent2_def,equiv_def,refl_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1076
diff changeset
   210
    "[| equiv(A,r);     \
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   211
\       !! 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
   212
\       !! 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
   213
\    |] ==> congruent2(r,b)";
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   214
by (cut_facts_tac prems 1);
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
   215
by Safe_tac;
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   216
by (rtac trans 1);
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   217
by (REPEAT (ares_tac prems 1
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   218
     ORELSE etac (subsetD RS SigmaE2) 1 THEN assume_tac 2 THEN assume_tac 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 749
diff changeset
   219
qed "congruent2I";
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   220
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5147
diff changeset
   221
val [equivA,commute,congt] = Goal
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1076
diff changeset
   222
    "[| equiv(A,r);     \
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   223
\       !! y z. [| y: A;  z: A |] ==> b(y,z) = b(z,y);        \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1076
diff changeset
   224
\       !! y z w. [| w: A;  <y,z>: r |] ==> b(w,y) = b(w,z)     \
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   225
\    |] ==> congruent2(r,b)";
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   226
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
   227
by (rtac (commute RS trans) 1);
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   228
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
   229
by (rtac sym 5);
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   230
by (REPEAT (ares_tac [congt] 1
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   231
     ORELSE etac (equivA RS equiv_type RS subsetD RS SigmaE2) 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 749
diff changeset
   232
qed "congruent2_commuteI";
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   233
723
82caba9e130f ZF/EquivClass/congruent_commuteI: uncommented and simplified proof
lcp
parents: 535
diff changeset
   234
(*Obsolete?*)
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5147
diff changeset
   235
val [equivA,ZinA,congt,commute] = Goalw [quotient_def]
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   236
    "[| equiv(A,r);  Z: A/r;  \
3840
e0baea4d485a fixed dots;
wenzelm
parents: 2929
diff changeset
   237
\       !!w. [| w: A |] ==> congruent(r, %z. b(w,z));    \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1076
diff changeset
   238
\       !!x y. [| x: A;  y: A |] ==> b(y,x) = b(x,y)    \
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   239
\    |] ==> 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
   240
val congt' = rewrite_rule [congruent_def] congt;
723
82caba9e130f ZF/EquivClass/congruent_commuteI: uncommented and simplified proof
lcp
parents: 535
diff changeset
   241
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
   242
by (rewtac congruent_def);
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
   243
by Safe_tac;
535
9d62c7e08699 ZF/ex/Equiv.thy,.ML: moved to main directory as EquivClass.thy,.ML
lcp
parents:
diff changeset
   244
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
   245
by (assume_tac 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3840
diff changeset
   246
by (asm_simp_tac (simpset() addsimps [commute,
2493
bdeb5024353a Removal of sum_cs and eq_cs
paulson
parents: 2469
diff changeset
   247
                                     [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
   248
by (REPEAT (ares_tac [congt' RS spec RS spec RS mp] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 749
diff changeset
   249
qed "congruent_commuteI";