| author | paulson | 
| Thu, 10 Jun 1999 10:35:58 +0200 | |
| changeset 6806 | 43c081a0858d | 
| parent 6005 | 45186ec4d8b6 | 
| child 7007 | b46ccfee8e59 | 
| 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  | 
|
| 5608 | 11  | 
Goalw [Id_def] "(a,a) : Id";  | 
| 2891 | 12  | 
by (Blast_tac 1);  | 
| 5608 | 13  | 
qed "IdI";  | 
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
14  | 
|
| 5608 | 15  | 
val major::prems = Goalw [Id_def]  | 
16  | 
"[| p: Id; !!x.[| p = (x,x) |] ==> P \  | 
|
| 
1128
 
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);  | 
| 5608 | 21  | 
qed "IdE";  | 
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
22  | 
|
| 5608 | 23  | 
Goalw [Id_def] "(a,b):Id = (a=b)";  | 
| 2891 | 24  | 
by (Blast_tac 1);  | 
| 5608 | 25  | 
qed "pair_in_Id_conv";  | 
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  | 
|
| 
6806
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
28  | 
Goalw [refl_def] "reflexive Id";  | 
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
29  | 
by Auto_tac;  | 
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
30  | 
qed "reflexive_Id";  | 
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
31  | 
|
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
32  | 
(*A strange result, since Id is also symmetric.*)  | 
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
33  | 
Goalw [antisym_def] "antisym Id";  | 
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
34  | 
by Auto_tac;  | 
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
35  | 
qed "antisym_Id";  | 
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
36  | 
|
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
37  | 
Goalw [trans_def] "trans Id";  | 
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
38  | 
by Auto_tac;  | 
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
39  | 
qed "trans_Id";  | 
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
40  | 
|
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
41  | 
|
| 
5978
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
42  | 
(** Diagonal relation: indentity restricted to some set **)  | 
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
43  | 
|
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
44  | 
(*** Equality : the diagonal relation ***)  | 
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
45  | 
|
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
46  | 
Goalw [diag_def] "[| a=b; a:A |] ==> (a,b) : diag(A)";  | 
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
47  | 
by (Blast_tac 1);  | 
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
48  | 
qed "diag_eqI";  | 
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
49  | 
|
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
50  | 
val diagI = refl RS diag_eqI |> standard;  | 
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
51  | 
|
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
52  | 
(*The general elimination rule*)  | 
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
53  | 
val major::prems = Goalw [diag_def]  | 
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
54  | 
"[| c : diag(A); \  | 
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
55  | 
\ !!x y. [| x:A; c = (x,x) |] ==> P \  | 
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
56  | 
\ |] ==> P";  | 
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
57  | 
by (rtac (major RS UN_E) 1);  | 
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
58  | 
by (REPEAT (eresolve_tac [asm_rl,singletonE] 1 ORELSE resolve_tac prems 1));  | 
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
59  | 
qed "diagE";  | 
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
60  | 
|
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
61  | 
AddSIs [diagI];  | 
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
62  | 
AddSEs [diagE];  | 
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
63  | 
|
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
64  | 
Goal "((x,y) : diag A) = (x=y & x : A)";  | 
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
65  | 
by (Blast_tac 1);  | 
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
66  | 
qed "diag_iff";  | 
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
67  | 
|
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
68  | 
Goal "diag(A) <= A Times A";  | 
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
69  | 
by (Blast_tac 1);  | 
| 5995 | 70  | 
qed "diag_subset_Times";  | 
| 
5978
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
71  | 
|
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
72  | 
|
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
73  | 
|
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
74  | 
(** Composition of two relations **)  | 
| 
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
75  | 
|
| 5069 | 76  | 
Goalw [comp_def]  | 
| 
5148
 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5143 
diff
changeset
 | 
77  | 
"[| (a,b):s; (b,c):r |] ==> (a,c) : r O s";  | 
| 2891 | 78  | 
by (Blast_tac 1);  | 
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
79  | 
qed "compI";  | 
| 
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
80  | 
|
| 
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
81  | 
(*proof requires higher-level assumptions or a delaying of hyp_subst_tac*)  | 
| 5316 | 82  | 
val prems = Goalw [comp_def]  | 
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
83  | 
"[| xz : r O s; \  | 
| 
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
84  | 
\ !!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
 | 
85  | 
\ |] ==> P";  | 
| 
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
86  | 
by (cut_facts_tac prems 1);  | 
| 
1985
 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 
paulson 
parents: 
1842 
diff
changeset
 | 
87  | 
by (REPEAT (eresolve_tac [CollectE, splitE, exE, conjE] 1  | 
| 
 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 
paulson 
parents: 
1842 
diff
changeset
 | 
88  | 
ORELSE ares_tac prems 1));  | 
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
89  | 
qed "compE";  | 
| 
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
90  | 
|
| 5316 | 91  | 
val prems = Goal  | 
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
92  | 
"[| (a,c) : r O s; \  | 
| 
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
93  | 
\ !!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
 | 
94  | 
\ |] ==> P";  | 
| 
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
95  | 
by (rtac compE 1);  | 
| 
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
96  | 
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
 | 
97  | 
qed "compEpair";  | 
| 
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
98  | 
|
| 5608 | 99  | 
AddIs [compI, IdI];  | 
100  | 
AddSEs [compE, IdE];  | 
|
| 
1754
 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 
berghofe 
parents: 
1694 
diff
changeset
 | 
101  | 
|
| 5608 | 102  | 
Goal "R O Id = R";  | 
| 4673 | 103  | 
by (Fast_tac 1);  | 
| 5608 | 104  | 
qed "R_O_Id";  | 
| 4673 | 105  | 
|
| 5608 | 106  | 
Goal "Id O R = R";  | 
| 4673 | 107  | 
by (Fast_tac 1);  | 
| 5608 | 108  | 
qed "Id_O_R";  | 
| 4673 | 109  | 
|
| 5608 | 110  | 
Addsimps [R_O_Id,Id_O_R];  | 
| 4673 | 111  | 
|
| 5069 | 112  | 
Goal "(R O S) O T = R O (S O T)";  | 
| 4830 | 113  | 
by (Blast_tac 1);  | 
114  | 
qed "O_assoc";  | 
|
115  | 
||
| 
5143
 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 
paulson 
parents: 
5069 
diff
changeset
 | 
116  | 
Goal "[| r'<=r; s'<=s |] ==> (r' O s') <= (r O s)";  | 
| 2891 | 117  | 
by (Blast_tac 1);  | 
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
118  | 
qed "comp_mono";  | 
| 
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
119  | 
|
| 
5148
 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5143 
diff
changeset
 | 
120  | 
Goal "[| s <= A Times B; r <= B Times C |] ==> (r O s) <= A Times C";  | 
| 2891 | 121  | 
by (Blast_tac 1);  | 
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
122  | 
qed "comp_subset_Sigma";  | 
| 
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
123  | 
|
| 
6806
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
124  | 
(** Natural deduction for refl(r) **)  | 
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
125  | 
|
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
126  | 
val prems = Goalw [refl_def]  | 
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
127  | 
"[| r <= A Times A; !! x. x:A ==> (x,x):r |] ==> refl A r";  | 
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
128  | 
by (REPEAT (ares_tac (prems@[ballI,conjI]) 1));  | 
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
129  | 
qed "reflI";  | 
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
130  | 
|
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
131  | 
Goalw [refl_def] "[| refl A r; a:A |] ==> (a,a):r";  | 
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
132  | 
by (Blast_tac 1);  | 
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
133  | 
qed "reflD";  | 
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
134  | 
|
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
135  | 
(** Natural deduction for antisym(r) **)  | 
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
136  | 
|
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
137  | 
val prems = Goalw [antisym_def]  | 
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
138  | 
"(!! x y. [| (x,y):r; (y,x):r |] ==> x=y) ==> antisym(r)";  | 
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
139  | 
by (REPEAT (ares_tac (prems@[allI,impI]) 1));  | 
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
140  | 
qed "antisymI";  | 
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
141  | 
|
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
142  | 
Goalw [antisym_def] "[| antisym(r); (a,b):r; (b,a):r |] ==> a=b";  | 
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
143  | 
by (Blast_tac 1);  | 
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
144  | 
qed "antisymD";  | 
| 
 
43c081a0858d
new preficates refl, sym [from Integ/Equiv], antisym
 
paulson 
parents: 
6005 
diff
changeset
 | 
145  | 
|
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
146  | 
(** Natural deduction for trans(r) **)  | 
| 
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
147  | 
|
| 5316 | 148  | 
val prems = Goalw [trans_def]  | 
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
149  | 
"(!! 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
 | 
150  | 
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
 | 
151  | 
qed "transI";  | 
| 
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
152  | 
|
| 
5148
 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5143 
diff
changeset
 | 
153  | 
Goalw [trans_def] "[| trans(r); (a,b):r; (b,c):r |] ==> (a,c):r";  | 
| 2891 | 154  | 
by (Blast_tac 1);  | 
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
155  | 
qed "transD";  | 
| 
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
156  | 
|
| 3439 | 157  | 
(** Natural deduction for r^-1 **)  | 
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
158  | 
|
| 
5143
 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 
paulson 
parents: 
5069 
diff
changeset
 | 
159  | 
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
 | 
160  | 
by (Simp_tac 1);  | 
| 4746 | 161  | 
qed "converse_iff";  | 
| 
1985
 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 
paulson 
parents: 
1842 
diff
changeset
 | 
162  | 
|
| 4746 | 163  | 
AddIffs [converse_iff];  | 
| 
1985
 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 
paulson 
parents: 
1842 
diff
changeset
 | 
164  | 
|
| 
5143
 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 
paulson 
parents: 
5069 
diff
changeset
 | 
165  | 
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
 | 
166  | 
by (Simp_tac 1);  | 
| 4746 | 167  | 
qed "converseI";  | 
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
168  | 
|
| 
5143
 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 
paulson 
parents: 
5069 
diff
changeset
 | 
169  | 
Goalw [converse_def] "(a,b) : r^-1 ==> (b,a) : r";  | 
| 2891 | 170  | 
by (Blast_tac 1);  | 
| 4746 | 171  | 
qed "converseD";  | 
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
172  | 
|
| 4746 | 173  | 
(*More general than converseD, as it "splits" the member of the relation*)  | 
174  | 
qed_goalw "converseE" thy [converse_def]  | 
|
| 3439 | 175  | 
"[| yx : r^-1; \  | 
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
176  | 
\ !!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
 | 
177  | 
\ |] ==> P"  | 
| 
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
178  | 
(fn [major,minor]=>  | 
| 
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
179  | 
[ (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
 | 
180  | 
(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
 | 
181  | 
(assume_tac 1) ]);  | 
| 
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
182  | 
|
| 4746 | 183  | 
AddSEs [converseE];  | 
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
184  | 
|
| 5069 | 185  | 
Goalw [converse_def] "(r^-1)^-1 = r";  | 
| 2891 | 186  | 
by (Blast_tac 1);  | 
| 4746 | 187  | 
qed "converse_converse";  | 
188  | 
Addsimps [converse_converse];  | 
|
| 
3413
 
c1f63cc3a768
Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
 
nipkow 
parents: 
2891 
diff
changeset
 | 
189  | 
|
| 5069 | 190  | 
Goal "(r O s)^-1 = s^-1 O r^-1";  | 
| 4423 | 191  | 
by (Blast_tac 1);  | 
| 4746 | 192  | 
qed "converse_comp";  | 
| 1605 | 193  | 
|
| 5608 | 194  | 
Goal "Id^-1 = Id";  | 
| 4644 | 195  | 
by (Blast_tac 1);  | 
| 5608 | 196  | 
qed "converse_Id";  | 
197  | 
Addsimps [converse_Id];  | 
|
| 4644 | 198  | 
|
| 5995 | 199  | 
Goal "(diag A) ^-1 = diag A";  | 
200  | 
by (Blast_tac 1);  | 
|
201  | 
qed "converse_diag";  | 
|
202  | 
Addsimps [converse_diag];  | 
|
203  | 
||
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
204  | 
(** Domain **)  | 
| 
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
205  | 
|
| 5811 | 206  | 
Goalw [Domain_def] "a: Domain(r) = (EX y. (a,y): r)";  | 
207  | 
by (Blast_tac 1);  | 
|
208  | 
qed "Domain_iff";  | 
|
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
209  | 
|
| 4673 | 210  | 
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
 | 
211  | 
(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
 | 
212  | 
|
| 4673 | 213  | 
qed_goal "DomainE" thy  | 
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
214  | 
"[| 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
 | 
215  | 
(fn prems=>  | 
| 
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
216  | 
[ (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
 | 
217  | 
(REPEAT (ares_tac prems 1)) ]);  | 
| 
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
218  | 
|
| 
1985
 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 
paulson 
parents: 
1842 
diff
changeset
 | 
219  | 
AddIs [DomainI];  | 
| 
 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 
paulson 
parents: 
1842 
diff
changeset
 | 
220  | 
AddSEs [DomainE];  | 
| 
 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 
paulson 
parents: 
1842 
diff
changeset
 | 
221  | 
|
| 5608 | 222  | 
Goal "Domain Id = UNIV";  | 
| 4644 | 223  | 
by (Blast_tac 1);  | 
| 5608 | 224  | 
qed "Domain_Id";  | 
225  | 
Addsimps [Domain_Id];  | 
|
| 4644 | 226  | 
|
| 
5978
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
227  | 
Goal "Domain (diag A) = A";  | 
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
228  | 
by Auto_tac;  | 
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
229  | 
qed "Domain_diag";  | 
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
230  | 
Addsimps [Domain_diag];  | 
| 
 
fa2c2dd74f8c
moved diag (diagonal relation) from Univ to Relation
 
paulson 
parents: 
5811 
diff
changeset
 | 
231  | 
|
| 5811 | 232  | 
Goal "Domain(A Un B) = Domain(A) Un Domain(B)";  | 
233  | 
by (Blast_tac 1);  | 
|
234  | 
qed "Domain_Un_eq";  | 
|
235  | 
||
236  | 
Goal "Domain(A Int B) <= Domain(A) Int Domain(B)";  | 
|
237  | 
by (Blast_tac 1);  | 
|
238  | 
qed "Domain_Int_subset";  | 
|
239  | 
||
240  | 
Goal "Domain(A) - Domain(B) <= Domain(A - B)";  | 
|
241  | 
by (Blast_tac 1);  | 
|
242  | 
qed "Domain_Diff_subset";  | 
|
243  | 
||
| 6005 | 244  | 
Goal "Domain (Union S) = (UN A:S. Domain A)";  | 
245  | 
by (Blast_tac 1);  | 
|
246  | 
qed "Domain_Union";  | 
|
247  | 
||
| 5811 | 248  | 
|
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
249  | 
(** Range **)  | 
| 
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
250  | 
|
| 5811 | 251  | 
Goalw [Domain_def, Range_def] "a: Range(r) = (EX y. (y,a): r)";  | 
252  | 
by (Blast_tac 1);  | 
|
253  | 
qed "Range_iff";  | 
|
254  | 
||
| 4673 | 255  | 
qed_goalw "RangeI" thy [Range_def] "!!a b r.(a,b): r ==> b : Range(r)"  | 
| 4746 | 256  | 
(fn _ => [ (etac (converseI RS DomainI) 1) ]);  | 
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
257  | 
|
| 4673 | 258  | 
qed_goalw "RangeE" thy [Range_def]  | 
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
259  | 
"[| 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
 | 
260  | 
(fn major::prems=>  | 
| 
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
261  | 
[ (rtac (major RS DomainE) 1),  | 
| 
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
262  | 
(resolve_tac prems 1),  | 
| 4746 | 263  | 
(etac converseD 1) ]);  | 
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
264  | 
|
| 
1985
 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 
paulson 
parents: 
1842 
diff
changeset
 | 
265  | 
AddIs [RangeI];  | 
| 
 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 
paulson 
parents: 
1842 
diff
changeset
 | 
266  | 
AddSEs [RangeE];  | 
| 
 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 
paulson 
parents: 
1842 
diff
changeset
 | 
267  | 
|
| 5608 | 268  | 
Goal "Range Id = UNIV";  | 
| 4644 | 269  | 
by (Blast_tac 1);  | 
| 5608 | 270  | 
qed "Range_Id";  | 
271  | 
Addsimps [Range_Id];  | 
|
| 4644 | 272  | 
|
| 5995 | 273  | 
Goal "Range (diag A) = A";  | 
274  | 
by Auto_tac;  | 
|
275  | 
qed "Range_diag";  | 
|
276  | 
Addsimps [Range_diag];  | 
|
277  | 
||
| 5811 | 278  | 
Goal "Range(A Un B) = Range(A) Un Range(B)";  | 
279  | 
by (Blast_tac 1);  | 
|
280  | 
qed "Range_Un_eq";  | 
|
281  | 
||
282  | 
Goal "Range(A Int B) <= Range(A) Int Range(B)";  | 
|
283  | 
by (Blast_tac 1);  | 
|
284  | 
qed "Range_Int_subset";  | 
|
285  | 
||
286  | 
Goal "Range(A) - Range(B) <= Range(A - B)";  | 
|
287  | 
by (Blast_tac 1);  | 
|
288  | 
qed "Range_Diff_subset";  | 
|
289  | 
||
| 6005 | 290  | 
Goal "Range (Union S) = (UN A:S. Range A)";  | 
291  | 
by (Blast_tac 1);  | 
|
292  | 
qed "Range_Union";  | 
|
293  | 
||
| 5811 | 294  | 
|
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
295  | 
(*** Image of a set under a relation ***)  | 
| 
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
296  | 
|
| 5649 | 297  | 
overload_1st_set "Relation.op ^^";  | 
| 5335 | 298  | 
|
| 4673 | 299  | 
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
 | 
300  | 
"b : r^^A = (? x:A. (x,b):r)"  | 
| 2891 | 301  | 
(fn _ => [ Blast_tac 1 ]);  | 
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
302  | 
|
| 4673 | 303  | 
qed_goalw "Image_singleton" thy [Image_def]  | 
304  | 
    "r^^{a} = {b. (a,b):r}"
 | 
|
305  | 
(fn _ => [ Blast_tac 1 ]);  | 
|
306  | 
||
307  | 
qed_goal "Image_singleton_iff" thy  | 
|
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
308  | 
    "(b : r^^{a}) = ((a,b):r)"
 | 
| 
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
309  | 
(fn _ => [ rtac (Image_iff RS trans) 1,  | 
| 2891 | 310  | 
Blast_tac 1 ]);  | 
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
311  | 
|
| 4673 | 312  | 
AddIffs [Image_singleton_iff];  | 
313  | 
||
314  | 
qed_goalw "ImageI" thy [Image_def]  | 
|
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
315  | 
"!!a b r. [| (a,b): r; a:A |] ==> b : r^^A"  | 
| 2891 | 316  | 
(fn _ => [ (Blast_tac 1)]);  | 
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
317  | 
|
| 4673 | 318  | 
qed_goalw "ImageE" thy [Image_def]  | 
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
319  | 
"[| 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
 | 
320  | 
(fn major::prems=>  | 
| 
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
321  | 
[ (rtac (major RS CollectE) 1),  | 
| 3718 | 322  | 
(Clarify_tac 1),  | 
| 
1128
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
323  | 
(rtac (hd prems) 1),  | 
| 
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
324  | 
(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
 | 
325  | 
|
| 
1985
 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 
paulson 
parents: 
1842 
diff
changeset
 | 
326  | 
AddIs [ImageI];  | 
| 
 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 
paulson 
parents: 
1842 
diff
changeset
 | 
327  | 
AddSEs [ImageE];  | 
| 
 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 
paulson 
parents: 
1842 
diff
changeset
 | 
328  | 
|
| 4593 | 329  | 
|
| 4673 | 330  | 
qed_goal "Image_empty" thy  | 
| 4593 | 331  | 
    "R^^{} = {}"
 | 
332  | 
(fn _ => [ Blast_tac 1 ]);  | 
|
333  | 
||
334  | 
Addsimps [Image_empty];  | 
|
335  | 
||
| 5608 | 336  | 
Goal "Id ^^ A = A";  | 
| 4601 | 337  | 
by (Blast_tac 1);  | 
| 5608 | 338  | 
qed "Image_Id";  | 
| 4601 | 339  | 
|
| 5998 | 340  | 
Goal "diag A ^^ B = A Int B";  | 
| 5995 | 341  | 
by (Blast_tac 1);  | 
342  | 
qed "Image_diag";  | 
|
343  | 
||
344  | 
Addsimps [Image_Id, Image_diag];  | 
|
| 4601 | 345  | 
|
| 4673 | 346  | 
qed_goal "Image_Int_subset" thy  | 
| 4593 | 347  | 
"R ^^ (A Int B) <= R ^^ A Int R ^^ B"  | 
348  | 
(fn _ => [ Blast_tac 1 ]);  | 
|
349  | 
||
| 
4733
 
2c984ac036f5
New theorem Image_eq_UN; deleted the silly vimage_inverse_Image
 
paulson 
parents: 
4673 
diff
changeset
 | 
350  | 
qed_goal "Image_Un" thy "R ^^ (A Un B) = R ^^ A Un R ^^ B"  | 
| 4593 | 351  | 
(fn _ => [ Blast_tac 1 ]);  | 
352  | 
||
| 
4733
 
2c984ac036f5
New theorem Image_eq_UN; deleted the silly vimage_inverse_Image
 
paulson 
parents: 
4673 
diff
changeset
 | 
353  | 
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
 | 
354  | 
(fn _ =>  | 
| 
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
355  | 
[ (rtac subsetI 1),  | 
| 
 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 
nipkow 
parents:  
diff
changeset
 | 
356  | 
(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
 | 
357  | 
|
| 
4733
 
2c984ac036f5
New theorem Image_eq_UN; deleted the silly vimage_inverse_Image
 
paulson 
parents: 
4673 
diff
changeset
 | 
358  | 
(*NOT suitable for rewriting*)  | 
| 5069 | 359  | 
Goal "r^^B = (UN y: B. r^^{y})";
 | 
| 4673 | 360  | 
by (Blast_tac 1);  | 
| 
4733
 
2c984ac036f5
New theorem Image_eq_UN; deleted the silly vimage_inverse_Image
 
paulson 
parents: 
4673 
diff
changeset
 | 
361  | 
qed "Image_eq_UN";  | 
| 
4760
 
9cdbd5a1d25a
added introduction and elimination rules for Univalent
 
oheimb 
parents: 
4746 
diff
changeset
 | 
362  | 
|
| 
 
9cdbd5a1d25a
added introduction and elimination rules for Univalent
 
oheimb 
parents: 
4746 
diff
changeset
 | 
363  | 
|
| 
 
9cdbd5a1d25a
added introduction and elimination rules for Univalent
 
oheimb 
parents: 
4746 
diff
changeset
 | 
364  | 
section "Univalent";  | 
| 
 
9cdbd5a1d25a
added introduction and elimination rules for Univalent
 
oheimb 
parents: 
4746 
diff
changeset
 | 
365  | 
|
| 
 
9cdbd5a1d25a
added introduction and elimination rules for Univalent
 
oheimb 
parents: 
4746 
diff
changeset
 | 
366  | 
qed_goalw "UnivalentI" Relation.thy [Univalent_def]  | 
| 
 
9cdbd5a1d25a
added introduction and elimination rules for Univalent
 
oheimb 
parents: 
4746 
diff
changeset
 | 
367  | 
"!!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
 | 
368  | 
|
| 
 
9cdbd5a1d25a
added introduction and elimination rules for Univalent
 
oheimb 
parents: 
4746 
diff
changeset
 | 
369  | 
qed_goalw "UnivalentD" Relation.thy [Univalent_def]  | 
| 
 
9cdbd5a1d25a
added introduction and elimination rules for Univalent
 
oheimb 
parents: 
4746 
diff
changeset
 | 
370  | 
"!!r. [| Univalent r; (x,y):r; (x,z):r|] ==> y=z" (K [Auto_tac]);  | 
| 5231 | 371  | 
|
372  | 
||
373  | 
(** Graphs of partial functions **)  | 
|
374  | 
||
375  | 
Goal "Domain{(x,y). y = f x & P x} = {x. P x}";
 | 
|
376  | 
by (Blast_tac 1);  | 
|
377  | 
qed "Domain_partial_func";  | 
|
378  | 
||
379  | 
Goal "Range{(x,y). y = f x & P x} = f``{x. P x}";
 | 
|
380  | 
by (Blast_tac 1);  | 
|
381  | 
qed "Range_partial_func";  | 
|
382  |