author | paulson |
Tue, 23 Jan 2001 15:46:25 +0100 | |
changeset 10963 | f2c1a280f1e3 |
parent 10864 | f0b0a125ae4b |
child 11080 | 22855d091249 |
permissions | -rw-r--r-- |
10341 | 1 |
(* ID: $Id$ *) |
10294 | 2 |
theory Relations = Main: |
3 |
||
4 |
ML "Pretty.setmargin 64" |
|
5 |
||
6 |
(*Id is only used in UNITY*) |
|
7 |
(*refl, antisym,trans,univalent,\<dots> ho hum*) |
|
8 |
||
9 |
text{* |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
10 |
@{thm[display] Id_def[no_vars]} |
10294 | 11 |
\rulename{Id_def} |
12 |
*} |
|
13 |
||
14 |
text{* |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
15 |
@{thm[display] comp_def[no_vars]} |
10294 | 16 |
\rulename{comp_def} |
17 |
*} |
|
18 |
||
19 |
text{* |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
20 |
@{thm[display] R_O_Id[no_vars]} |
10294 | 21 |
\rulename{R_O_Id} |
22 |
*} |
|
23 |
||
24 |
text{* |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
25 |
@{thm[display] comp_mono[no_vars]} |
10294 | 26 |
\rulename{comp_mono} |
27 |
*} |
|
28 |
||
29 |
text{* |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
30 |
@{thm[display] converse_iff[no_vars]} |
10294 | 31 |
\rulename{converse_iff} |
32 |
*} |
|
33 |
||
34 |
text{* |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
35 |
@{thm[display] converse_comp[no_vars]} |
10294 | 36 |
\rulename{converse_comp} |
37 |
*} |
|
38 |
||
39 |
text{* |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
40 |
@{thm[display] Image_iff[no_vars]} |
10294 | 41 |
\rulename{Image_iff} |
42 |
*} |
|
43 |
||
44 |
text{* |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
45 |
@{thm[display] Image_UN[no_vars]} |
10294 | 46 |
\rulename{Image_UN} |
47 |
*} |
|
48 |
||
49 |
text{* |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
50 |
@{thm[display] Domain_iff[no_vars]} |
10294 | 51 |
\rulename{Domain_iff} |
52 |
*} |
|
53 |
||
54 |
text{* |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
55 |
@{thm[display] Range_iff[no_vars]} |
10294 | 56 |
\rulename{Range_iff} |
57 |
*} |
|
58 |
||
59 |
text{* |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
60 |
@{thm[display] relpow.simps[no_vars]} |
10294 | 61 |
\rulename{relpow.simps} |
62 |
||
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
63 |
@{thm[display] rtrancl_unfold[no_vars]} |
10294 | 64 |
\rulename{rtrancl_unfold} |
65 |
||
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
66 |
@{thm[display] rtrancl_refl[no_vars]} |
10294 | 67 |
\rulename{rtrancl_refl} |
68 |
||
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
69 |
@{thm[display] r_into_rtrancl[no_vars]} |
10294 | 70 |
\rulename{r_into_rtrancl} |
71 |
||
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
72 |
@{thm[display] rtrancl_trans[no_vars]} |
10294 | 73 |
\rulename{rtrancl_trans} |
74 |
||
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
75 |
@{thm[display] rtrancl_induct[no_vars]} |
10294 | 76 |
\rulename{rtrancl_induct} |
77 |
||
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
78 |
@{thm[display] rtrancl_idemp[no_vars]} |
10294 | 79 |
\rulename{rtrancl_idemp} |
80 |
||
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
81 |
@{thm[display] r_into_trancl[no_vars]} |
10294 | 82 |
\rulename{r_into_trancl} |
83 |
||
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
84 |
@{thm[display] trancl_trans[no_vars]} |
10294 | 85 |
\rulename{trancl_trans} |
86 |
||
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
87 |
@{thm[display] trancl_into_rtrancl[no_vars]} |
10294 | 88 |
\rulename{trancl_into_rtrancl} |
89 |
||
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
90 |
@{thm[display] trancl_converse[no_vars]} |
10294 | 91 |
\rulename{trancl_converse} |
92 |
*} |
|
93 |
||
94 |
text{*Relations. transitive closure*} |
|
95 |
||
96 |
lemma rtrancl_converseD: "(x,y) \<in> (r^-1)^* \<Longrightarrow> (y,x) \<in> r^*" |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
97 |
apply (erule rtrancl_induct) |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
98 |
txt{* |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
99 |
@{subgoals[display,indent=0,margin=65]} |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
100 |
*}; |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
101 |
apply (rule rtrancl_refl) |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
102 |
apply (blast intro: r_into_rtrancl rtrancl_trans) |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
103 |
done |
10294 | 104 |
|
105 |
||
106 |
lemma rtrancl_converseI: "(y,x) \<in> r^* \<Longrightarrow> (x,y) \<in> (r^-1)^*" |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
107 |
apply (erule rtrancl_induct) |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
108 |
apply (rule rtrancl_refl) |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
109 |
apply (blast intro: r_into_rtrancl rtrancl_trans) |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
110 |
done |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
111 |
|
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
112 |
lemma rtrancl_converse: "(r^-1)^* = (r^*)^-1" |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
113 |
by (auto intro: rtrancl_converseI dest: rtrancl_converseD) |
10294 | 114 |
|
115 |
lemma rtrancl_converse: "(r^-1)^* = (r^*)^-1" |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
116 |
apply (intro equalityI subsetI) |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
117 |
txt{* |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
118 |
after intro rules |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
119 |
|
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
120 |
@{subgoals[display,indent=0,margin=65]} |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
121 |
*}; |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
122 |
apply clarify |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
123 |
txt{* |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
124 |
after splitting |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
125 |
@{subgoals[display,indent=0,margin=65]} |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
126 |
*}; |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
127 |
oops |
10294 | 128 |
|
129 |
||
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
130 |
lemma "(\<forall>u v. (u,v) \<in> A \<longrightarrow> u=v) \<Longrightarrow> A \<subseteq> Id" |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
131 |
apply (rule subsetI) |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
132 |
txt{* |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
133 |
@{subgoals[display,indent=0,margin=65]} |
10294 | 134 |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
135 |
after subsetI |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
136 |
*}; |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
137 |
apply clarify |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
138 |
txt{* |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
139 |
@{subgoals[display,indent=0,margin=65]} |
10294 | 140 |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
141 |
subgoals after clarify |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
142 |
*}; |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
143 |
by blast |
10294 | 144 |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
145 |
|
10294 | 146 |
|
147 |
||
148 |
text{*rejects*} |
|
149 |
||
150 |
lemma "(a \<in> {z. P z} \<union> {y. Q y}) = P a \<or> Q a" |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
151 |
apply (blast) |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
152 |
done |
10294 | 153 |
|
154 |
text{*Pow, Inter too little used*} |
|
155 |
||
156 |
lemma "(A \<subset> B) = (A \<subseteq> B \<and> A \<noteq> B)" |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
157 |
apply (simp add: psubset_def) |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
158 |
done |
10294 | 159 |
|
160 |
end |