| author | nipkow | 
| Fri, 29 Nov 1996 15:07:27 +0100 | |
| changeset 2279 | 2f337bf81085 | 
| parent 2215 | ebf910e7ec87 | 
| child 3358 | 13f1df323daf | 
| permissions | -rw-r--r-- | 
| 1465 | 1 | (* Title: Equiv.ML | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 2 | ID: $Id$ | 
| 2215 | 3 | Authors: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 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 | 9 | val RSLIST = curry (op MRS); | 
| 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 | 13 | Delrules [equalityI]; | 
| 14 | ||
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 15 | (*** Suppes, Theorem 70: r is an equiv relation iff converse(r) O r = r ***) | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 16 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 17 | (** first half: equiv A r ==> converse(r) O r = r **) | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 18 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 19 | goalw Equiv.thy [trans_def,sym_def,converse_def] | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 20 | "!!r. [| sym(r); trans(r) |] ==> converse(r) O r <= r"; | 
| 1894 | 21 | by (fast_tac (!claset addSEs [converseD]) 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 | 24 | goalw Equiv.thy [refl_def] | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 25 | "!!A r. refl A r ==> r <= converse(r) O r"; | 
| 1894 | 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] | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 30 | "!!A r. equiv A r ==> converse(r) O r = r"; | 
| 
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] | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 38 | "!!A r. [| converse(r) O r = r; Domain(r) = A |] ==> equiv A r"; | 
| 
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: 
925diff
changeset | 40 | by (subgoal_tac "ALL x y. (x,y) : r --> (y,x) : r" 1); | 
| 2215 | 41 | by (Step_tac 1); | 
| 1894 | 42 | by (fast_tac (!claset addSIs [converseI] addIs [compI]) 3); | 
| 43 | by (ALLGOALS (fast_tac (!claset addIs [compI] addSEs [compE]))); | |
| 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: 
925diff
changeset | 50 |     "!!A r. [| equiv A r;  (a,b): r |] ==> r^^{a} <= r^^{b}";
 | 
| 2215 | 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 | 53 | by (Fast_tac 2); | 
| 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: 
925diff
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 | 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 | 63 | goalw Equiv.thy [equiv_def,refl_def] | 
| 64 |     "!!A r. [| equiv A r;  a: A |] ==> a: r^^{a}";
 | |
| 1894 | 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: 
925diff
changeset | 70 |     "!!A r. [| equiv A r;  r^^{b} <= r^^{a};  b: A |] ==> (a,b): r";
 | 
| 1894 | 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 | 74 | goal Equiv.thy | 
| 75 |     "!!A r. [| r^^{a} = r^^{b};  equiv A r;  b: A |] ==> (a,b): r";
 | |
| 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: 
925diff
changeset | 81 |     "!!A r. [| equiv A r;  x: (r^^{a} Int r^^{b}) |] ==> (a,b): r";
 | 
| 1894 | 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 | 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: 
925diff
changeset | 91 |     "!!A r. equiv A r ==> ((x,y): r) = (r^^{x} = r^^{y} & x:A & y:A)";
 | 
| 2215 | 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: 
925diff
changeset | 103 |     "!!A r. [| equiv A r;  x: A;  y: A |] ==> (r^^{x} = r^^{y}) = ((x,y): r)";
 | 
| 2215 | 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 | 115 | goalw Equiv.thy [quotient_def] "!!A. x:A ==> r^^{x}: A/r";
 | 
| 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 | 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 | 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 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 128 | (** Not needed by Theory Integ --> bypassed **) | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 129 | (**goalw Equiv.thy [equiv_def,refl_def,quotient_def] | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 130 | "!!A r. equiv A r ==> Union(A/r) = A"; | 
| 1844 | 131 | by (Fast_tac 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 132 | qed "Union_quotient"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 133 | **) | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 134 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 135 | (** Not needed by Theory Integ --> bypassed **) | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 136 | (*goalw Equiv.thy [quotient_def] | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 137 | "!!A r. [| equiv A r; X: A/r; Y: A/r |] ==> X=Y | (X Int Y <= 0)"; | 
| 1894 | 138 | by (safe_tac (!claset addSIs [equiv_class_eq])); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 139 | by (assume_tac 1); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 140 | by (rewrite_goals_tac [equiv_def,trans_def,sym_def]); | 
| 1894 | 141 | by (Fast_tac 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 142 | qed "quotient_disj"; | 
| 
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 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 145 | (**** Defining unary operations upon equivalence classes ****) | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 146 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 147 | (* theorem needed to prove UN_equiv_class *) | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 148 | goal Set.thy "!!A. [| a:A; ! y:A. b(y)=b(a) |] ==> (UN y:A. b(y))=b(a)"; | 
| 1894 | 149 | by (fast_tac (!claset addSEs [equalityE] addSIs [equalityI]) 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 150 | qed "UN_singleton_lemma"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 151 | val UN_singleton = ballI RSN (2,UN_singleton_lemma); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 152 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 153 | |
| 2215 | 154 | (** These proofs really require the local premises | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 155 | equiv A r; congruent r b | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 156 | **) | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 157 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 158 | (*Conversion rule*) | 
| 2215 | 159 | goal Equiv.thy "!!A r. [| equiv A r; congruent r b; a: A |] \ | 
| 160 | \                      ==> (UN x:r^^{a}. b(x)) = b(a)";
 | |
| 161 | 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 | 162 | by (rewrite_goals_tac [equiv_def,congruent_def,sym_def]); | 
| 1894 | 163 | by (Fast_tac 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 164 | qed "UN_equiv_class"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 165 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 166 | (*type checking of  UN x:r``{a}. b(x) *)
 | 
| 2215 | 167 | val prems = goalw Equiv.thy [quotient_def] | 
| 1465 | 168 | "[| equiv A r; congruent r b; X: A/r; \ | 
| 2215 | 169 | \ !!x. x : A ==> b(x) : B |] \ | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 170 | \ ==> (UN x:X. b(x)) : B"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 171 | by (cut_facts_tac prems 1); | 
| 2215 | 172 | by (Step_tac 1); | 
| 173 | by (stac UN_equiv_class 1); | |
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 174 | by (REPEAT (ares_tac prems 1)); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 175 | qed "UN_equiv_class_type"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 176 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 177 | (*Sufficient conditions for injectiveness. Could weaken premises! | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 178 | 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 | 179 | *) | 
| 2215 | 180 | val prems = goalw Equiv.thy [quotient_def] | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 181 | "[| equiv A r; congruent r b; \ | 
| 2215 | 182 | \ (UN x:X. b(x))=(UN y:Y. b(y)); X: A/r; Y: A/r; \ | 
| 1465 | 183 | \ !!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 | 184 | \ ==> X=Y"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 185 | by (cut_facts_tac prems 1); | 
| 2215 | 186 | by (Step_tac 1); | 
| 187 | by (rtac equiv_class_eq 1); | |
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 188 | by (REPEAT (ares_tac prems 1)); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 189 | by (etac box_equals 1); | 
| 2215 | 190 | by (REPEAT (ares_tac [UN_equiv_class] 1)); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 191 | qed "UN_equiv_class_inject"; | 
| 
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 | (**** Defining binary operations upon equivalence classes ****) | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 195 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 196 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 197 | 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 | 198 | "!!A r. [| equiv A r; congruent2 r b; a: A |] ==> congruent r (b a)"; | 
| 1894 | 199 | by (Fast_tac 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 200 | qed "congruent2_implies_congruent"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 201 | |
| 2215 | 202 | goalw Equiv.thy [congruent_def] | 
| 203 | "!!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 | 204 | \    congruent r (%x1. UN x2:r^^{a}. b x1 x2)";
 | 
| 2215 | 205 | by (Step_tac 1); | 
| 206 | by (rtac (equiv_type RS subsetD RS SigmaE2) 1 THEN REPEAT (assume_tac 1)); | |
| 207 | by (asm_simp_tac (!simpset addsimps [UN_equiv_class, | |
| 1465 | 208 | congruent2_implies_congruent]) 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 209 | by (rewrite_goals_tac [congruent2_def,equiv_def,refl_def]); | 
| 1894 | 210 | by (Fast_tac 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 211 | qed "congruent2_implies_congruent_UN"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 212 | |
| 2215 | 213 | goal Equiv.thy | 
| 214 | "!!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 | 215 | \    ==> (UN x1:r^^{a1}. UN x2:r^^{a2}. b x1 x2) = b a1 a2";
 | 
| 2215 | 216 | by (asm_simp_tac (!simpset addsimps [UN_equiv_class, | 
| 1465 | 217 | congruent2_implies_congruent, | 
| 218 | congruent2_implies_congruent_UN]) 1); | |
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 219 | qed "UN_equiv_class2"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 220 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 221 | (*type checking*) | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 222 | val prems = goalw Equiv.thy [quotient_def] | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 223 | "[| equiv A r; congruent2 r b; \ | 
| 1465 | 224 | \ X1: A/r; X2: A/r; \ | 
| 225 | \ !!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 | 226 | \ ==> (UN x1:X1. UN x2:X2. b x1 x2) : B"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 227 | by (cut_facts_tac prems 1); | 
| 2215 | 228 | by (Step_tac 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 229 | by (REPEAT (ares_tac (prems@[UN_equiv_class_type, | 
| 1465 | 230 | congruent2_implies_congruent_UN, | 
| 231 | congruent2_implies_congruent, quotientI]) 1)); | |
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 232 | qed "UN_equiv_class_type2"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 233 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 234 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 235 | (*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 | 236 | than the direct proof*) | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 237 | val prems = goalw Equiv.thy [congruent2_def,equiv_def,refl_def] | 
| 1465 | 238 | "[| equiv A r; \ | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
925diff
changeset | 239 | \ !! y z w. [| w: A; (y,z) : r |] ==> b y w = b z w; \ | 
| 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
925diff
changeset | 240 | \ !! 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 | 241 | \ |] ==> congruent2 r b"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 242 | by (cut_facts_tac prems 1); | 
| 2215 | 243 | by (Step_tac 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 244 | by (rtac trans 1); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 245 | by (REPEAT (ares_tac prems 1 | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 246 | 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 | 247 | qed "congruent2I"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 248 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 249 | val [equivA,commute,congt] = goal Equiv.thy | 
| 1465 | 250 | "[| equiv A r; \ | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 251 | \ !! y z. [| y: A; z: A |] ==> b y z = b z y; \ | 
| 1465 | 252 | \ !! 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 | 253 | \ |] ==> congruent2 r b"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 254 | by (resolve_tac [equivA RS congruent2I] 1); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 255 | by (rtac (commute RS trans) 1); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 256 | by (rtac (commute RS trans RS sym) 3); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 257 | by (rtac sym 5); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 258 | by (REPEAT (ares_tac [congt] 1 | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 259 | 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 | 260 | qed "congruent2_commuteI"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 261 |