author | paulson |
Wed, 15 Jul 1998 10:15:13 +0200 | |
changeset 5143 | b94cd208f073 |
parent 5069 | 3ea049f7979d |
child 5148 | 74919e8f221c |
permissions | -rw-r--r-- |
1465 | 1 |
(* Title: Relation.ML |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
2 |
ID: $Id$ |
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
1842
diff
changeset
|
3 |
Authors: Lawrence C Paulson, Cambridge University Computer Laboratory |
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
1842
diff
changeset
|
4 |
Copyright 1996 University of Cambridge |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
5 |
*) |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
6 |
|
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
7 |
open Relation; |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
8 |
|
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
9 |
(** Identity relation **) |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
10 |
|
5069 | 11 |
Goalw [id_def] "(a,a) : id"; |
2891 | 12 |
by (Blast_tac 1); |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
13 |
qed "idI"; |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
14 |
|
4673 | 15 |
val major::prems = goalw thy [id_def] |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
16 |
"[| p: id; !!x.[| p = (x,x) |] ==> P \ |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
17 |
\ |] ==> P"; |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
18 |
by (rtac (major RS CollectE) 1); |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
19 |
by (etac exE 1); |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
20 |
by (eresolve_tac prems 1); |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
21 |
qed "idE"; |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
22 |
|
5069 | 23 |
Goalw [id_def] "(a,b):id = (a=b)"; |
2891 | 24 |
by (Blast_tac 1); |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
25 |
qed "pair_in_id_conv"; |
1694 | 26 |
Addsimps [pair_in_id_conv]; |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
27 |
|
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
28 |
|
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
29 |
(** Composition of two relations **) |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
30 |
|
5069 | 31 |
Goalw [comp_def] |
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
1842
diff
changeset
|
32 |
"!!r s. [| (a,b):s; (b,c):r |] ==> (a,c) : r O s"; |
2891 | 33 |
by (Blast_tac 1); |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
34 |
qed "compI"; |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
35 |
|
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
36 |
(*proof requires higher-level assumptions or a delaying of hyp_subst_tac*) |
4673 | 37 |
val prems = goalw thy [comp_def] |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
38 |
"[| xz : r O s; \ |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
39 |
\ !!x y z. [| xz = (x,z); (x,y):s; (y,z):r |] ==> P \ |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
40 |
\ |] ==> P"; |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
41 |
by (cut_facts_tac prems 1); |
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
1842
diff
changeset
|
42 |
by (REPEAT (eresolve_tac [CollectE, splitE, exE, conjE] 1 |
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
1842
diff
changeset
|
43 |
ORELSE ares_tac prems 1)); |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
44 |
qed "compE"; |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
45 |
|
4673 | 46 |
val prems = goal thy |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
47 |
"[| (a,c) : r O s; \ |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
48 |
\ !!y. [| (a,y):s; (y,c):r |] ==> P \ |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
49 |
\ |] ==> P"; |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
50 |
by (rtac compE 1); |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
51 |
by (REPEAT (ares_tac prems 1 ORELSE eresolve_tac [Pair_inject,ssubst] 1)); |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
52 |
qed "compEpair"; |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
53 |
|
1754
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents:
1694
diff
changeset
|
54 |
AddIs [compI, idI]; |
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents:
1694
diff
changeset
|
55 |
AddSEs [compE, idE]; |
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents:
1694
diff
changeset
|
56 |
|
5069 | 57 |
Goal "R O id = R"; |
4673 | 58 |
by (Fast_tac 1); |
59 |
qed "R_O_id"; |
|
60 |
||
5069 | 61 |
Goal "id O R = R"; |
4673 | 62 |
by (Fast_tac 1); |
63 |
qed "id_O_R"; |
|
64 |
||
65 |
Addsimps [R_O_id,id_O_R]; |
|
66 |
||
5069 | 67 |
Goal "(R O S) O T = R O (S O T)"; |
4830 | 68 |
by (Blast_tac 1); |
69 |
qed "O_assoc"; |
|
70 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
71 |
Goal "[| r'<=r; s'<=s |] ==> (r' O s') <= (r O s)"; |
2891 | 72 |
by (Blast_tac 1); |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
73 |
qed "comp_mono"; |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
74 |
|
5069 | 75 |
Goal |
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
1842
diff
changeset
|
76 |
"!!r s. [| s <= A Times B; r <= B Times C |] ==> (r O s) <= A Times C"; |
2891 | 77 |
by (Blast_tac 1); |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
78 |
qed "comp_subset_Sigma"; |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
79 |
|
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
80 |
(** Natural deduction for trans(r) **) |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
81 |
|
4673 | 82 |
val prems = goalw thy [trans_def] |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
83 |
"(!! x y z. [| (x,y):r; (y,z):r |] ==> (x,z):r) ==> trans(r)"; |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
84 |
by (REPEAT (ares_tac (prems@[allI,impI]) 1)); |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
85 |
qed "transI"; |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
86 |
|
5069 | 87 |
Goalw [trans_def] |
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
1842
diff
changeset
|
88 |
"!!r. [| trans(r); (a,b):r; (b,c):r |] ==> (a,c):r"; |
2891 | 89 |
by (Blast_tac 1); |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
90 |
qed "transD"; |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
91 |
|
3439 | 92 |
(** Natural deduction for r^-1 **) |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
93 |
|
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
94 |
Goalw [converse_def] "((a,b): r^-1) = ((b,a):r)"; |
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
1842
diff
changeset
|
95 |
by (Simp_tac 1); |
4746 | 96 |
qed "converse_iff"; |
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
1842
diff
changeset
|
97 |
|
4746 | 98 |
AddIffs [converse_iff]; |
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
1842
diff
changeset
|
99 |
|
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
100 |
Goalw [converse_def] "(a,b):r ==> (b,a): r^-1"; |
1264
3eb91524b938
added local simpsets; removed IOA from 'make test'
clasohm
parents:
1128
diff
changeset
|
101 |
by (Simp_tac 1); |
4746 | 102 |
qed "converseI"; |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
103 |
|
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5069
diff
changeset
|
104 |
Goalw [converse_def] "(a,b) : r^-1 ==> (b,a) : r"; |
2891 | 105 |
by (Blast_tac 1); |
4746 | 106 |
qed "converseD"; |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
107 |
|
4746 | 108 |
(*More general than converseD, as it "splits" the member of the relation*) |
109 |
qed_goalw "converseE" thy [converse_def] |
|
3439 | 110 |
"[| yx : r^-1; \ |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
111 |
\ !!x y. [| yx=(y,x); (x,y):r |] ==> P \ |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
112 |
\ |] ==> P" |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
113 |
(fn [major,minor]=> |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
114 |
[ (rtac (major RS CollectE) 1), |
1454
d0266c81a85e
Streamlined defs in Relation and added new intro/elim rules to do with
nipkow
parents:
1264
diff
changeset
|
115 |
(REPEAT (eresolve_tac [splitE, bexE,exE, conjE, minor] 1)), |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
116 |
(assume_tac 1) ]); |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
117 |
|
4746 | 118 |
AddSEs [converseE]; |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
119 |
|
5069 | 120 |
Goalw [converse_def] "(r^-1)^-1 = r"; |
2891 | 121 |
by (Blast_tac 1); |
4746 | 122 |
qed "converse_converse"; |
123 |
Addsimps [converse_converse]; |
|
3413
c1f63cc3a768
Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
nipkow
parents:
2891
diff
changeset
|
124 |
|
5069 | 125 |
Goal "(r O s)^-1 = s^-1 O r^-1"; |
4423 | 126 |
by (Blast_tac 1); |
4746 | 127 |
qed "converse_comp"; |
1605 | 128 |
|
5069 | 129 |
Goal "id^-1 = id"; |
4644 | 130 |
by (Blast_tac 1); |
4746 | 131 |
qed "converse_id"; |
132 |
Addsimps [converse_id]; |
|
4644 | 133 |
|
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
134 |
(** Domain **) |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
135 |
|
4673 | 136 |
qed_goalw "Domain_iff" thy [Domain_def] |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
137 |
"a: Domain(r) = (EX y. (a,y): r)" |
2891 | 138 |
(fn _=> [ (Blast_tac 1) ]); |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
139 |
|
4673 | 140 |
qed_goal "DomainI" thy "!!a b r. (a,b): r ==> a: Domain(r)" |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
141 |
(fn _ => [ (etac (exI RS (Domain_iff RS iffD2)) 1) ]); |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
142 |
|
4673 | 143 |
qed_goal "DomainE" thy |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
144 |
"[| a : Domain(r); !!y. (a,y): r ==> P |] ==> P" |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
145 |
(fn prems=> |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
146 |
[ (rtac (Domain_iff RS iffD1 RS exE) 1), |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
147 |
(REPEAT (ares_tac prems 1)) ]); |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
148 |
|
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
1842
diff
changeset
|
149 |
AddIs [DomainI]; |
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
1842
diff
changeset
|
150 |
AddSEs [DomainE]; |
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
1842
diff
changeset
|
151 |
|
5069 | 152 |
Goal "Domain id = UNIV"; |
4644 | 153 |
by (Blast_tac 1); |
154 |
qed "Domain_id"; |
|
155 |
Addsimps [Domain_id]; |
|
156 |
||
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
157 |
(** Range **) |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
158 |
|
4673 | 159 |
qed_goalw "RangeI" thy [Range_def] "!!a b r.(a,b): r ==> b : Range(r)" |
4746 | 160 |
(fn _ => [ (etac (converseI RS DomainI) 1) ]); |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
161 |
|
4673 | 162 |
qed_goalw "RangeE" thy [Range_def] |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
163 |
"[| b : Range(r); !!x. (x,b): r ==> P |] ==> P" |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
164 |
(fn major::prems=> |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
165 |
[ (rtac (major RS DomainE) 1), |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
166 |
(resolve_tac prems 1), |
4746 | 167 |
(etac converseD 1) ]); |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
168 |
|
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
1842
diff
changeset
|
169 |
AddIs [RangeI]; |
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
1842
diff
changeset
|
170 |
AddSEs [RangeE]; |
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
1842
diff
changeset
|
171 |
|
5069 | 172 |
Goal "Range id = UNIV"; |
4644 | 173 |
by (Blast_tac 1); |
174 |
qed "Range_id"; |
|
175 |
Addsimps [Range_id]; |
|
176 |
||
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
177 |
(*** Image of a set under a relation ***) |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
178 |
|
4673 | 179 |
qed_goalw "Image_iff" thy [Image_def] |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
180 |
"b : r^^A = (? x:A. (x,b):r)" |
2891 | 181 |
(fn _ => [ Blast_tac 1 ]); |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
182 |
|
4673 | 183 |
qed_goalw "Image_singleton" thy [Image_def] |
184 |
"r^^{a} = {b. (a,b):r}" |
|
185 |
(fn _ => [ Blast_tac 1 ]); |
|
186 |
||
187 |
qed_goal "Image_singleton_iff" thy |
|
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
188 |
"(b : r^^{a}) = ((a,b):r)" |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
189 |
(fn _ => [ rtac (Image_iff RS trans) 1, |
2891 | 190 |
Blast_tac 1 ]); |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
191 |
|
4673 | 192 |
AddIffs [Image_singleton_iff]; |
193 |
||
194 |
qed_goalw "ImageI" thy [Image_def] |
|
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
195 |
"!!a b r. [| (a,b): r; a:A |] ==> b : r^^A" |
2891 | 196 |
(fn _ => [ (Blast_tac 1)]); |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
197 |
|
4673 | 198 |
qed_goalw "ImageE" thy [Image_def] |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
199 |
"[| b: r^^A; !!x.[| (x,b): r; x:A |] ==> P |] ==> P" |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
200 |
(fn major::prems=> |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
201 |
[ (rtac (major RS CollectE) 1), |
3718 | 202 |
(Clarify_tac 1), |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
203 |
(rtac (hd prems) 1), |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
204 |
(REPEAT (etac bexE 1 ORELSE ares_tac prems 1)) ]); |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
205 |
|
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
1842
diff
changeset
|
206 |
AddIs [ImageI]; |
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
1842
diff
changeset
|
207 |
AddSEs [ImageE]; |
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
1842
diff
changeset
|
208 |
|
4593 | 209 |
|
4673 | 210 |
qed_goal "Image_empty" thy |
4593 | 211 |
"R^^{} = {}" |
212 |
(fn _ => [ Blast_tac 1 ]); |
|
213 |
||
214 |
Addsimps [Image_empty]; |
|
215 |
||
5069 | 216 |
Goal "id ^^ A = A"; |
4601 | 217 |
by (Blast_tac 1); |
218 |
qed "Image_id"; |
|
219 |
||
220 |
Addsimps [Image_id]; |
|
221 |
||
4673 | 222 |
qed_goal "Image_Int_subset" thy |
4593 | 223 |
"R ^^ (A Int B) <= R ^^ A Int R ^^ B" |
224 |
(fn _ => [ Blast_tac 1 ]); |
|
225 |
||
4733
2c984ac036f5
New theorem Image_eq_UN; deleted the silly vimage_inverse_Image
paulson
parents:
4673
diff
changeset
|
226 |
qed_goal "Image_Un" thy "R ^^ (A Un B) = R ^^ A Un R ^^ B" |
4593 | 227 |
(fn _ => [ Blast_tac 1 ]); |
228 |
||
4733
2c984ac036f5
New theorem Image_eq_UN; deleted the silly vimage_inverse_Image
paulson
parents:
4673
diff
changeset
|
229 |
qed_goal "Image_subset" thy "!!A B r. r <= A Times B ==> r^^C <= B" |
1128
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
230 |
(fn _ => |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
231 |
[ (rtac subsetI 1), |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
232 |
(REPEAT (eresolve_tac [asm_rl, ImageE, subsetD RS SigmaD2] 1)) ]); |
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
nipkow
parents:
diff
changeset
|
233 |
|
4733
2c984ac036f5
New theorem Image_eq_UN; deleted the silly vimage_inverse_Image
paulson
parents:
4673
diff
changeset
|
234 |
(*NOT suitable for rewriting*) |
5069 | 235 |
Goal "r^^B = (UN y: B. r^^{y})"; |
4673 | 236 |
by (Blast_tac 1); |
4733
2c984ac036f5
New theorem Image_eq_UN; deleted the silly vimage_inverse_Image
paulson
parents:
4673
diff
changeset
|
237 |
qed "Image_eq_UN"; |
4760
9cdbd5a1d25a
added introduction and elimination rules for Univalent
oheimb
parents:
4746
diff
changeset
|
238 |
|
9cdbd5a1d25a
added introduction and elimination rules for Univalent
oheimb
parents:
4746
diff
changeset
|
239 |
|
9cdbd5a1d25a
added introduction and elimination rules for Univalent
oheimb
parents:
4746
diff
changeset
|
240 |
section "Univalent"; |
9cdbd5a1d25a
added introduction and elimination rules for Univalent
oheimb
parents:
4746
diff
changeset
|
241 |
|
9cdbd5a1d25a
added introduction and elimination rules for Univalent
oheimb
parents:
4746
diff
changeset
|
242 |
qed_goalw "UnivalentI" Relation.thy [Univalent_def] |
9cdbd5a1d25a
added introduction and elimination rules for Univalent
oheimb
parents:
4746
diff
changeset
|
243 |
"!!r. !x y. (x,y):r --> (!z. (x,z):r --> y=z) ==> Univalent r" (K [atac 1]); |
9cdbd5a1d25a
added introduction and elimination rules for Univalent
oheimb
parents:
4746
diff
changeset
|
244 |
|
9cdbd5a1d25a
added introduction and elimination rules for Univalent
oheimb
parents:
4746
diff
changeset
|
245 |
qed_goalw "UnivalentD" Relation.thy [Univalent_def] |
9cdbd5a1d25a
added introduction and elimination rules for Univalent
oheimb
parents:
4746
diff
changeset
|
246 |
"!!r. [| Univalent r; (x,y):r; (x,z):r|] ==> y=z" (K [Auto_tac]); |