src/HOL/Integ/Equiv.ML
author nipkow
Tue, 17 Jun 1997 09:01:56 +0200
changeset 3439 54785105178c
parent 3427 e7cef2081106
child 3457 a8ab7c64817c
permissions -rw-r--r--
converse -> ^-1
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
     1
(*  Title:      Equiv.ML
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
2215
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
     3
    Authors:    Lawrence C Paulson, Cambridge University Computer Laboratory
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
     4
    Copyright   1996  University of Cambridge
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
     5
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
     6
Equivalence relations in HOL Set Theory 
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
     7
*)
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
     8
1978
e7df069acb74 Moved RSLIST here from ../Relation.ML
paulson
parents: 1894
diff changeset
     9
val RSLIST = curry (op MRS);
e7df069acb74 Moved RSLIST here from ../Relation.ML
paulson
parents: 1894
diff changeset
    10
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    11
open Equiv;
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    12
1894
c2c8279d40f0 Classical tactics now use default claset.
berghofe
parents: 1844
diff changeset
    13
Delrules [equalityI];
c2c8279d40f0 Classical tactics now use default claset.
berghofe
parents: 1844
diff changeset
    14
3439
54785105178c converse -> ^-1
nipkow
parents: 3427
diff changeset
    15
(*** Suppes, Theorem 70: r is an equiv relation iff r^-1 O r = r ***)
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    16
3439
54785105178c converse -> ^-1
nipkow
parents: 3427
diff changeset
    17
(** first half: equiv A r ==> r^-1 O r = r **)
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    18
3439
54785105178c converse -> ^-1
nipkow
parents: 3427
diff changeset
    19
goalw Equiv.thy [trans_def,sym_def,inverse_def]
54785105178c converse -> ^-1
nipkow
parents: 3427
diff changeset
    20
    "!!r. [| sym(r); trans(r) |] ==> r^-1 O r <= r";
54785105178c converse -> ^-1
nipkow
parents: 3427
diff changeset
    21
by (fast_tac (!claset addSEs [inverseD]) 1);
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    22
qed "sym_trans_comp_subset";
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    23
1045
0cdf86973c49 Deleted some useless things and made proofs of
lcp
parents: 972
diff changeset
    24
goalw Equiv.thy [refl_def]
3439
54785105178c converse -> ^-1
nipkow
parents: 3427
diff changeset
    25
    "!!A r. refl A r ==> r <= r^-1 O r";
1894
c2c8279d40f0 Classical tactics now use default claset.
berghofe
parents: 1844
diff changeset
    26
by (fast_tac (!claset addIs [compI]) 1);
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    27
qed "refl_comp_subset";
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    28
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    29
goalw Equiv.thy [equiv_def]
3439
54785105178c converse -> ^-1
nipkow
parents: 3427
diff changeset
    30
    "!!A r. equiv A r ==> r^-1 O r = r";
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    31
by (rtac equalityI 1);
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    32
by (REPEAT (ares_tac [sym_trans_comp_subset, refl_comp_subset] 1
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    33
     ORELSE etac conjE 1));
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    34
qed "equiv_comp_eq";
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    35
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    36
(*second half*)
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    37
goalw Equiv.thy [equiv_def,refl_def,sym_def,trans_def]
3439
54785105178c converse -> ^-1
nipkow
parents: 3427
diff changeset
    38
    "!!A r. [| r^-1 O r = r;  Domain(r) = A |] ==> equiv A r";
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    39
by (etac equalityE 1);
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 925
diff changeset
    40
by (subgoal_tac "ALL x y. (x,y) : r --> (y,x) : r" 1);
2215
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
    41
by (Step_tac 1);
3427
e7cef2081106 Removed a few redundant additions of simprules or classical rules
paulson
parents: 3374
diff changeset
    42
by (fast_tac (!claset addIs [compI]) 3);
e7cef2081106 Removed a few redundant additions of simprules or classical rules
paulson
parents: 3374
diff changeset
    43
by (ALLGOALS (fast_tac (!claset addIs [compI])));
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    44
qed "comp_equivI";
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    45
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    46
(** Equivalence classes **)
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    47
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    48
(*Lemma for the next result*)
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    49
goalw Equiv.thy [equiv_def,trans_def,sym_def]
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 925
diff changeset
    50
    "!!A r. [| equiv A r;  (a,b): r |] ==> r^^{a} <= r^^{b}";
2215
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
    51
by (Step_tac 1);
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    52
by (rtac ImageI 1);
1894
c2c8279d40f0 Classical tactics now use default claset.
berghofe
parents: 1844
diff changeset
    53
by (Fast_tac 2);
c2c8279d40f0 Classical tactics now use default claset.
berghofe
parents: 1844
diff changeset
    54
by (Fast_tac 1);
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    55
qed "equiv_class_subset";
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    56
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 925
diff changeset
    57
goal Equiv.thy "!!A r. [| equiv A r;  (a,b): r |] ==> r^^{a} = r^^{b}";
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    58
by (REPEAT (ares_tac [equalityI, equiv_class_subset] 1));
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    59
by (rewrite_goals_tac [equiv_def,sym_def]);
1894
c2c8279d40f0 Classical tactics now use default claset.
berghofe
parents: 1844
diff changeset
    60
by (Fast_tac 1);
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    61
qed "equiv_class_eq";
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    62
2215
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
    63
goalw Equiv.thy [equiv_def,refl_def]
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
    64
    "!!A r. [| equiv A r;  a: A |] ==> a: r^^{a}";
1894
c2c8279d40f0 Classical tactics now use default claset.
berghofe
parents: 1844
diff changeset
    65
by (Fast_tac 1);
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    66
qed "equiv_class_self";
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    67
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    68
(*Lemma for the next result*)
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    69
goalw Equiv.thy [equiv_def,refl_def]
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 925
diff changeset
    70
    "!!A r. [| equiv A r;  r^^{b} <= r^^{a};  b: A |] ==> (a,b): r";
1894
c2c8279d40f0 Classical tactics now use default claset.
berghofe
parents: 1844
diff changeset
    71
by (Fast_tac 1);
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    72
qed "subset_equiv_class";
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    73
2215
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
    74
goal Equiv.thy
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
    75
    "!!A r. [| r^^{a} = r^^{b};  equiv A r;  b: A |] ==> (a,b): r";
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
    76
by (REPEAT (ares_tac [equalityD2, subset_equiv_class] 1));
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    77
qed "eq_equiv_class";
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    78
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    79
(*thus r^^{a} = r^^{b} as well*)
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    80
goalw Equiv.thy [equiv_def,trans_def,sym_def]
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 925
diff changeset
    81
    "!!A r. [| equiv A r;  x: (r^^{a} Int r^^{b}) |] ==> (a,b): r";
1894
c2c8279d40f0 Classical tactics now use default claset.
berghofe
parents: 1844
diff changeset
    82
by (Fast_tac 1);
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    83
qed "equiv_class_nondisjoint";
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    84
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    85
val [major] = goalw Equiv.thy [equiv_def,refl_def]
1642
21db0cf9a1a4 Using new "Times" infix
paulson
parents: 1465
diff changeset
    86
    "equiv A r ==> r <= A Times A";
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    87
by (rtac (major RS conjunct1 RS conjunct1) 1);
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    88
qed "equiv_type";
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    89
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    90
goal Equiv.thy
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 925
diff changeset
    91
    "!!A r. equiv A r ==> ((x,y): r) = (r^^{x} = r^^{y} & x:A & y:A)";
2215
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
    92
by (Step_tac 1);
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    93
by ((rtac equiv_class_eq 1) THEN (assume_tac 1) THEN (assume_tac 1));
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    94
by ((rtac eq_equiv_class 3) THEN 
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    95
    (assume_tac 4) THEN (assume_tac 4) THEN (assume_tac 3));
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    96
by ((dtac equiv_type 1) THEN (dtac rev_subsetD 1) THEN
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    97
    (assume_tac 1) THEN (dtac SigmaD1 1) THEN (assume_tac 1));
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    98
by ((dtac equiv_type 1) THEN (dtac rev_subsetD 1) THEN
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
    99
    (assume_tac 1) THEN (dtac SigmaD2 1) THEN (assume_tac 1));
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   100
qed "equiv_class_eq_iff";
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   101
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   102
goal Equiv.thy
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 925
diff changeset
   103
    "!!A r. [| equiv A r;  x: A;  y: A |] ==> (r^^{x} = r^^{y}) = ((x,y): r)";
2215
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
   104
by (Step_tac 1);
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   105
by ((rtac eq_equiv_class 1) THEN 
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   106
    (assume_tac 1) THEN (assume_tac 1) THEN (assume_tac 1));
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   107
by ((rtac equiv_class_eq 1) THEN 
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   108
    (assume_tac 1) THEN (assume_tac 1));
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   109
qed "eq_equiv_class_iff";
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   110
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   111
(*** Quotients ***)
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   112
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   113
(** Introduction/elimination rules -- needed? **)
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   114
2215
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
   115
goalw Equiv.thy [quotient_def] "!!A. x:A ==> r^^{x}: A/r";
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
   116
by (Fast_tac 1);
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   117
qed "quotientI";
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   118
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   119
val [major,minor] = goalw Equiv.thy [quotient_def]
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
   120
    "[| X:(A/r);  !!x. [| X = r^^{x};  x:A |] ==> P |]  \
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   121
\    ==> P";
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   122
by (resolve_tac [major RS UN_E] 1);
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   123
by (rtac minor 1);
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   124
by (assume_tac 2);
1894
c2c8279d40f0 Classical tactics now use default claset.
berghofe
parents: 1844
diff changeset
   125
by (Fast_tac 1);
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   126
qed "quotientE";
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   127
3358
13f1df323daf Restoration of the two "bypassed" theorems Union_quotient and quotient_disj
paulson
parents: 2215
diff changeset
   128
goalw Equiv.thy [equiv_def,refl_def,quotient_def]
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   129
    "!!A r. equiv A r ==> Union(A/r) = A";
3358
13f1df323daf Restoration of the two "bypassed" theorems Union_quotient and quotient_disj
paulson
parents: 2215
diff changeset
   130
by (blast_tac (!claset addSIs [equalityI]) 1);
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   131
qed "Union_quotient";
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   132
3358
13f1df323daf Restoration of the two "bypassed" theorems Union_quotient and quotient_disj
paulson
parents: 2215
diff changeset
   133
goalw Equiv.thy [quotient_def]
13f1df323daf Restoration of the two "bypassed" theorems Union_quotient and quotient_disj
paulson
parents: 2215
diff changeset
   134
    "!!A r. [| equiv A r;  X: A/r;  Y: A/r |] ==> X=Y | (X Int Y = {})";
1894
c2c8279d40f0 Classical tactics now use default claset.
berghofe
parents: 1844
diff changeset
   135
by (safe_tac (!claset addSIs [equiv_class_eq]));
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   136
by (assume_tac 1);
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   137
by (rewrite_goals_tac [equiv_def,trans_def,sym_def]);
3358
13f1df323daf Restoration of the two "bypassed" theorems Union_quotient and quotient_disj
paulson
parents: 2215
diff changeset
   138
by (blast_tac (!claset addSIs [equalityI]) 1);
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   139
qed "quotient_disj";
3358
13f1df323daf Restoration of the two "bypassed" theorems Union_quotient and quotient_disj
paulson
parents: 2215
diff changeset
   140
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   141
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   142
(**** Defining unary operations upon equivalence classes ****)
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   143
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   144
(* theorem needed to prove UN_equiv_class *)
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   145
goal Set.thy "!!A. [| a:A; ! y:A. b(y)=b(a) |] ==> (UN y:A. b(y))=b(a)";
1894
c2c8279d40f0 Classical tactics now use default claset.
berghofe
parents: 1844
diff changeset
   146
by (fast_tac (!claset addSEs [equalityE] addSIs [equalityI]) 1);
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   147
qed "UN_singleton_lemma";
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   148
val UN_singleton = ballI RSN (2,UN_singleton_lemma);
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   149
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   150
2215
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
   151
(** These proofs really require the local premises
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   152
     equiv A r;  congruent r b
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   153
**)
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   154
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   155
(*Conversion rule*)
2215
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
   156
goal Equiv.thy "!!A r. [| equiv A r;  congruent r b;  a: A |] \
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
   157
\                      ==> (UN x:r^^{a}. b(x)) = b(a)";
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
   158
by (rtac (equiv_class_self RS UN_singleton) 1 THEN REPEAT (assume_tac 1));
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   159
by (rewrite_goals_tac [equiv_def,congruent_def,sym_def]);
1894
c2c8279d40f0 Classical tactics now use default claset.
berghofe
parents: 1844
diff changeset
   160
by (Fast_tac 1);
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   161
qed "UN_equiv_class";
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   162
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   163
(*type checking of  UN x:r``{a}. b(x) *)
2215
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
   164
val prems = goalw Equiv.thy [quotient_def]
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
   165
    "[| equiv A r;  congruent r b;  X: A/r;     \
2215
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
   166
\       !!x.  x : A ==> b(x) : B |]             \
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   167
\    ==> (UN x:X. b(x)) : B";
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   168
by (cut_facts_tac prems 1);
2215
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
   169
by (Step_tac 1);
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
   170
by (stac UN_equiv_class 1);
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   171
by (REPEAT (ares_tac prems 1));
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   172
qed "UN_equiv_class_type";
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   173
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   174
(*Sufficient conditions for injectiveness.  Could weaken premises!
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   175
  major premise could be an inclusion; bcong could be !!y. y:A ==> b(y):B
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   176
*)
2215
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
   177
val prems = goalw Equiv.thy [quotient_def]
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   178
    "[| equiv A r;   congruent r b;  \
2215
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
   179
\       (UN x:X. b(x))=(UN y:Y. b(y));  X: A/r;  Y: A/r;        \
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
   180
\       !!x y. [| x:A; y:A; b(x)=b(y) |] ==> (x,y):r |]         \
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   181
\    ==> X=Y";
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   182
by (cut_facts_tac prems 1);
2215
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
   183
by (Step_tac 1);
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
   184
by (rtac equiv_class_eq 1);
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   185
by (REPEAT (ares_tac prems 1));
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   186
by (etac box_equals 1);
2215
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
   187
by (REPEAT (ares_tac [UN_equiv_class] 1));
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   188
qed "UN_equiv_class_inject";
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   189
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   190
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   191
(**** Defining binary operations upon equivalence classes ****)
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   192
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   193
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   194
goalw Equiv.thy [congruent_def,congruent2_def,equiv_def,refl_def]
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   195
    "!!A r. [| equiv A r;  congruent2 r b;  a: A |] ==> congruent r (b a)";
1894
c2c8279d40f0 Classical tactics now use default claset.
berghofe
parents: 1844
diff changeset
   196
by (Fast_tac 1);
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   197
qed "congruent2_implies_congruent";
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   198
2215
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
   199
goalw Equiv.thy [congruent_def]
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
   200
    "!!A r. [| equiv A r;  congruent2 r b;  a: A |] ==> \
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   201
\    congruent r (%x1. UN x2:r^^{a}. b x1 x2)";
2215
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
   202
by (Step_tac 1);
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
   203
by (rtac (equiv_type RS subsetD RS SigmaE2) 1 THEN REPEAT (assume_tac 1));
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
   204
by (asm_simp_tac (!simpset addsimps [UN_equiv_class,
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
   205
                                     congruent2_implies_congruent]) 1);
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   206
by (rewrite_goals_tac [congruent2_def,equiv_def,refl_def]);
1894
c2c8279d40f0 Classical tactics now use default claset.
berghofe
parents: 1844
diff changeset
   207
by (Fast_tac 1);
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   208
qed "congruent2_implies_congruent_UN";
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   209
2215
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
   210
goal Equiv.thy
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
   211
    "!!A r. [| equiv A r;  congruent2 r b;  a1: A;  a2: A |]  \
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   212
\    ==> (UN x1:r^^{a1}. UN x2:r^^{a2}. b x1 x2) = b a1 a2";
2215
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
   213
by (asm_simp_tac (!simpset addsimps [UN_equiv_class,
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
   214
                                     congruent2_implies_congruent,
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
   215
                                     congruent2_implies_congruent_UN]) 1);
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   216
qed "UN_equiv_class2";
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   217
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   218
(*type checking*)
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   219
val prems = goalw Equiv.thy [quotient_def]
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   220
    "[| equiv A r;  congruent2 r b;  \
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
   221
\       X1: A/r;  X2: A/r;      \
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
   222
\       !!x1 x2.  [| x1: A; x2: A |] ==> b x1 x2 : B |]    \
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   223
\    ==> (UN x1:X1. UN x2:X2. b x1 x2) : B";
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   224
by (cut_facts_tac prems 1);
2215
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
   225
by (Step_tac 1);
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   226
by (REPEAT (ares_tac (prems@[UN_equiv_class_type,
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
   227
                             congruent2_implies_congruent_UN,
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
   228
                             congruent2_implies_congruent, quotientI]) 1));
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   229
qed "UN_equiv_class_type2";
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   230
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   231
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   232
(*Suggested by John Harrison -- the two subproofs may be MUCH simpler
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   233
  than the direct proof*)
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   234
val prems = goalw Equiv.thy [congruent2_def,equiv_def,refl_def]
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
   235
    "[| equiv A r;      \
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 925
diff changeset
   236
\       !! y z w. [| w: A;  (y,z) : r |] ==> b y w = b z w;      \
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 925
diff changeset
   237
\       !! y z w. [| w: A;  (y,z) : r |] ==> b w y = b w z       \
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   238
\    |] ==> congruent2 r b";
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   239
by (cut_facts_tac prems 1);
2215
ebf910e7ec87 Tidied up some proofs, ...
paulson
parents: 2036
diff changeset
   240
by (Step_tac 1);
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   241
by (rtac trans 1);
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   242
by (REPEAT (ares_tac prems 1
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   243
     ORELSE etac (subsetD RS SigmaE2) 1 THEN assume_tac 2 THEN assume_tac 1));
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   244
qed "congruent2I";
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   245
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   246
val [equivA,commute,congt] = goal Equiv.thy
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
   247
    "[| equiv A r;      \
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   248
\       !! y z. [| y: A;  z: A |] ==> b y z = b z y;        \
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1266
diff changeset
   249
\       !! y z w. [| w: A;  (y,z): r |] ==> b w y = b w z       \
925
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   250
\    |] ==> congruent2 r b";
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   251
by (resolve_tac [equivA RS congruent2I] 1);
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   252
by (rtac (commute RS trans) 1);
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   253
by (rtac (commute RS trans RS sym) 3);
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   254
by (rtac sym 5);
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   255
by (REPEAT (ares_tac [congt] 1
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   256
     ORELSE etac (equivA RS equiv_type RS subsetD RS SigmaE2) 1));
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   257
qed "congruent2_commuteI";
15539deb6863 new version of HOL/Integ with curried function application
clasohm
parents:
diff changeset
   258
3374
182a2b76d19e New proofs about cardinality. Suggested by Florian Kammueller
paulson
parents: 3358
diff changeset
   259
182a2b76d19e New proofs about cardinality. Suggested by Florian Kammueller
paulson
parents: 3358
diff changeset
   260
(*** Cardinality results suggested by Florian Kammüller ***)
182a2b76d19e New proofs about cardinality. Suggested by Florian Kammueller
paulson
parents: 3358
diff changeset
   261
182a2b76d19e New proofs about cardinality. Suggested by Florian Kammueller
paulson
parents: 3358
diff changeset
   262
(*Recall that  equiv A r  implies  r <= A Times A   (equiv_type) *)
182a2b76d19e New proofs about cardinality. Suggested by Florian Kammueller
paulson
parents: 3358
diff changeset
   263
goal thy "!!A. [| finite A; r <= A Times A |] ==> finite (A/r)";
182a2b76d19e New proofs about cardinality. Suggested by Florian Kammueller
paulson
parents: 3358
diff changeset
   264
by (rtac finite_subset 1);
182a2b76d19e New proofs about cardinality. Suggested by Florian Kammueller
paulson
parents: 3358
diff changeset
   265
by (etac (finite_Pow_iff RS iffD2) 2);
182a2b76d19e New proofs about cardinality. Suggested by Florian Kammueller
paulson
parents: 3358
diff changeset
   266
by (rewrite_goals_tac [quotient_def]);
182a2b76d19e New proofs about cardinality. Suggested by Florian Kammueller
paulson
parents: 3358
diff changeset
   267
by (Blast_tac 1);
182a2b76d19e New proofs about cardinality. Suggested by Florian Kammueller
paulson
parents: 3358
diff changeset
   268
qed "finite_quotient";
182a2b76d19e New proofs about cardinality. Suggested by Florian Kammueller
paulson
parents: 3358
diff changeset
   269
182a2b76d19e New proofs about cardinality. Suggested by Florian Kammueller
paulson
parents: 3358
diff changeset
   270
goalw thy [quotient_def]
182a2b76d19e New proofs about cardinality. Suggested by Florian Kammueller
paulson
parents: 3358
diff changeset
   271
    "!!A. [| finite A;  r <= A Times A;  X : A/r |] ==> finite X";
182a2b76d19e New proofs about cardinality. Suggested by Florian Kammueller
paulson
parents: 3358
diff changeset
   272
by (rtac finite_subset 1);
182a2b76d19e New proofs about cardinality. Suggested by Florian Kammueller
paulson
parents: 3358
diff changeset
   273
ba 2;
182a2b76d19e New proofs about cardinality. Suggested by Florian Kammueller
paulson
parents: 3358
diff changeset
   274
by (Blast_tac 1);
182a2b76d19e New proofs about cardinality. Suggested by Florian Kammueller
paulson
parents: 3358
diff changeset
   275
qed "finite_equiv_class";
182a2b76d19e New proofs about cardinality. Suggested by Florian Kammueller
paulson
parents: 3358
diff changeset
   276
182a2b76d19e New proofs about cardinality. Suggested by Florian Kammueller
paulson
parents: 3358
diff changeset
   277
goal thy "!!A. [| finite A;  equiv A r;  ! X : A/r. k dvd card(X) |] \
182a2b76d19e New proofs about cardinality. Suggested by Florian Kammueller
paulson
parents: 3358
diff changeset
   278
\              ==> k dvd card(A)";
182a2b76d19e New proofs about cardinality. Suggested by Florian Kammueller
paulson
parents: 3358
diff changeset
   279
by (rtac (Union_quotient RS subst) 1);
182a2b76d19e New proofs about cardinality. Suggested by Florian Kammueller
paulson
parents: 3358
diff changeset
   280
by (assume_tac 1);
182a2b76d19e New proofs about cardinality. Suggested by Florian Kammueller
paulson
parents: 3358
diff changeset
   281
by (rtac dvd_partition 1);
182a2b76d19e New proofs about cardinality. Suggested by Florian Kammueller
paulson
parents: 3358
diff changeset
   282
by (blast_tac (!claset delrules [equalityI] addEs [quotient_disj RS disjE]) 4);
182a2b76d19e New proofs about cardinality. Suggested by Florian Kammueller
paulson
parents: 3358
diff changeset
   283
by (ALLGOALS 
182a2b76d19e New proofs about cardinality. Suggested by Florian Kammueller
paulson
parents: 3358
diff changeset
   284
    (asm_simp_tac (!simpset addsimps [Union_quotient, equiv_type, 
182a2b76d19e New proofs about cardinality. Suggested by Florian Kammueller
paulson
parents: 3358
diff changeset
   285
				      finite_quotient])));
182a2b76d19e New proofs about cardinality. Suggested by Florian Kammueller
paulson
parents: 3358
diff changeset
   286
qed "equiv_imp_dvd_card";