| author | paulson | 
| Wed, 23 Jul 1997 11:52:22 +0200 | |
| changeset 3564 | f886dbd91ee5 | 
| parent 3439 | 54785105178c | 
| child 3718 | d78cf498a88c | 
| 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: 
1842diff
changeset | 3 | Authors: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1842diff
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 | |
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 11 | goalw Relation.thy [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 | |
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 15 | val major::prems = goalw Relation.thy [id_def] | 
| 
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 | |
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 23 | goalw Relation.thy [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 | |
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1842diff
changeset | 31 | goalw Relation.thy [comp_def] | 
| 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1842diff
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*) | 
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 37 | val prems = goalw Relation.thy [comp_def] | 
| 
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: 
1842diff
changeset | 42 | by (REPEAT (eresolve_tac [CollectE, splitE, exE, conjE] 1 | 
| 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1842diff
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 | |
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 46 | val prems = goal Relation.thy | 
| 
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: 
1694diff
changeset | 54 | AddIs [compI, idI]; | 
| 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 berghofe parents: 
1694diff
changeset | 55 | AddSEs [compE, idE]; | 
| 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 berghofe parents: 
1694diff
changeset | 56 | |
| 1128 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 57 | goal Relation.thy "!!r s. [| r'<=r; s'<=s |] ==> (r' O s') <= (r O s)"; | 
| 2891 | 58 | by (Blast_tac 1); | 
| 1128 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 59 | qed "comp_mono"; | 
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 60 | |
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 61 | goal Relation.thy | 
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1842diff
changeset | 62 | "!!r s. [| s <= A Times B; r <= B Times C |] ==> (r O s) <= A Times C"; | 
| 2891 | 63 | by (Blast_tac 1); | 
| 1128 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 64 | qed "comp_subset_Sigma"; | 
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 65 | |
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 66 | (** Natural deduction for trans(r) **) | 
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 67 | |
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 68 | val prems = goalw Relation.thy [trans_def] | 
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 69 | "(!! 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 | 70 | 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 | 71 | qed "transI"; | 
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 72 | |
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1842diff
changeset | 73 | goalw Relation.thy [trans_def] | 
| 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1842diff
changeset | 74 | "!!r. [| trans(r); (a,b):r; (b,c):r |] ==> (a,c):r"; | 
| 2891 | 75 | by (Blast_tac 1); | 
| 1128 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 76 | qed "transD"; | 
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 77 | |
| 3439 | 78 | (** Natural deduction for r^-1 **) | 
| 1128 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 79 | |
| 3439 | 80 | goalw Relation.thy [inverse_def] "!!a b r. ((a,b): r^-1) = ((b,a):r)"; | 
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1842diff
changeset | 81 | by (Simp_tac 1); | 
| 3439 | 82 | qed "inverse_iff"; | 
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1842diff
changeset | 83 | |
| 3439 | 84 | AddIffs [inverse_iff]; | 
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1842diff
changeset | 85 | |
| 3439 | 86 | goalw Relation.thy [inverse_def] "!!a b r. (a,b):r ==> (b,a): r^-1"; | 
| 1264 
3eb91524b938
added local simpsets; removed IOA from 'make test'
 clasohm parents: 
1128diff
changeset | 87 | by (Simp_tac 1); | 
| 3439 | 88 | qed "inverseI"; | 
| 1128 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 89 | |
| 3439 | 90 | goalw Relation.thy [inverse_def] "!!a b r. (a,b) : r^-1 ==> (b,a) : r"; | 
| 2891 | 91 | by (Blast_tac 1); | 
| 3439 | 92 | qed "inverseD"; | 
| 1128 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 93 | |
| 3439 | 94 | (*More general than inverseD, as it "splits" the member of the relation*) | 
| 95 | qed_goalw "inverseE" Relation.thy [inverse_def] | |
| 96 | "[| yx : r^-1; \ | |
| 1128 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 97 | \ !!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 | 98 | \ |] ==> P" | 
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 99 | (fn [major,minor]=> | 
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 100 | [ (rtac (major RS CollectE) 1), | 
| 1454 
d0266c81a85e
Streamlined defs in Relation and added new intro/elim rules to do with
 nipkow parents: 
1264diff
changeset | 101 | (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 | 102 | (assume_tac 1) ]); | 
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 103 | |
| 3439 | 104 | AddSEs [inverseE]; | 
| 1128 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 105 | |
| 3439 | 106 | goalw Relation.thy [inverse_def] "(r^-1)^-1 = r"; | 
| 2891 | 107 | by (Blast_tac 1); | 
| 3439 | 108 | qed "inverse_inverse"; | 
| 109 | Addsimps [inverse_inverse]; | |
| 3413 
c1f63cc3a768
Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
 nipkow parents: 
2891diff
changeset | 110 | |
| 3439 | 111 | goal Relation.thy "(r O s)^-1 = s^-1 O r^-1"; | 
| 3413 
c1f63cc3a768
Finite.ML Finite.thy: Replaced `finite subset of' by mere `finite'.
 nipkow parents: 
2891diff
changeset | 112 | by(Blast_tac 1); | 
| 3439 | 113 | qed "inverse_comp"; | 
| 1605 | 114 | |
| 1128 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 115 | (** Domain **) | 
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 116 | |
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 117 | qed_goalw "Domain_iff" Relation.thy [Domain_def] | 
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 118 | "a: Domain(r) = (EX y. (a,y): r)" | 
| 2891 | 119 | (fn _=> [ (Blast_tac 1) ]); | 
| 1128 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 120 | |
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 121 | qed_goal "DomainI" Relation.thy "!!a b r. (a,b): r ==> a: Domain(r)" | 
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 122 | (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 | 123 | |
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 124 | qed_goal "DomainE" Relation.thy | 
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 125 | "[| 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 | 126 | (fn prems=> | 
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 127 | [ (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 | 128 | (REPEAT (ares_tac prems 1)) ]); | 
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 129 | |
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1842diff
changeset | 130 | AddIs [DomainI]; | 
| 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1842diff
changeset | 131 | AddSEs [DomainE]; | 
| 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1842diff
changeset | 132 | |
| 1128 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 133 | (** Range **) | 
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 134 | |
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 135 | qed_goalw "RangeI" Relation.thy [Range_def] "!!a b r.(a,b): r ==> b : Range(r)" | 
| 3439 | 136 | (fn _ => [ (etac (inverseI RS DomainI) 1) ]); | 
| 1128 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 137 | |
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 138 | qed_goalw "RangeE" Relation.thy [Range_def] | 
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 139 | "[| 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 | 140 | (fn major::prems=> | 
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 141 | [ (rtac (major RS DomainE) 1), | 
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 142 | (resolve_tac prems 1), | 
| 3439 | 143 | (etac inverseD 1) ]); | 
| 1128 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 144 | |
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1842diff
changeset | 145 | AddIs [RangeI]; | 
| 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1842diff
changeset | 146 | AddSEs [RangeE]; | 
| 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1842diff
changeset | 147 | |
| 1128 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 148 | (*** Image of a set under a relation ***) | 
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 149 | |
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 150 | qed_goalw "Image_iff" Relation.thy [Image_def] | 
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 151 | "b : r^^A = (? x:A. (x,b):r)" | 
| 2891 | 152 | (fn _ => [ Blast_tac 1 ]); | 
| 1128 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 153 | |
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 154 | qed_goal "Image_singleton_iff" Relation.thy | 
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 155 |     "(b : r^^{a}) = ((a,b):r)"
 | 
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 156 | (fn _ => [ rtac (Image_iff RS trans) 1, | 
| 2891 | 157 | Blast_tac 1 ]); | 
| 1128 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 158 | |
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 159 | qed_goalw "ImageI" Relation.thy [Image_def] | 
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 160 | "!!a b r. [| (a,b): r; a:A |] ==> b : r^^A" | 
| 2891 | 161 | (fn _ => [ (Blast_tac 1)]); | 
| 1128 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 162 | |
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 163 | qed_goalw "ImageE" Relation.thy [Image_def] | 
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 164 | "[| 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 | 165 | (fn major::prems=> | 
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 166 | [ (rtac (major RS CollectE) 1), | 
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1842diff
changeset | 167 | (Step_tac 1), | 
| 1128 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 168 | (rtac (hd prems) 1), | 
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 169 | (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 | 170 | |
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1842diff
changeset | 171 | AddIs [ImageI]; | 
| 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1842diff
changeset | 172 | AddSEs [ImageE]; | 
| 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1842diff
changeset | 173 | |
| 1128 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 174 | qed_goal "Image_subset" Relation.thy | 
| 1642 | 175 | "!!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 | 176 | (fn _ => | 
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 177 | [ (rtac subsetI 1), | 
| 
64b30e3cc6d4
Trancl is now based on Relation which used to be in Integ.
 nipkow parents: diff
changeset | 178 | (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 | 179 | |
| 1694 | 180 | goal Relation.thy "R O id = R"; | 
| 2637 
e9b203f854ae
reflecting my recent changes of the simplifier and classical reasoner
 oheimb parents: 
2031diff
changeset | 181 | by (fast_tac (!claset addbefore split_all_tac) 1); | 
| 1694 | 182 | qed "R_O_id"; | 
| 183 | ||
| 184 | goal Relation.thy "id O R = R"; | |
| 2637 
e9b203f854ae
reflecting my recent changes of the simplifier and classical reasoner
 oheimb parents: 
2031diff
changeset | 185 | by (fast_tac (!claset addbefore split_all_tac) 1); | 
| 1694 | 186 | qed "id_O_R"; | 
| 187 | ||
| 188 | Addsimps [R_O_id,id_O_R]; |