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