author | webertj |
Fri, 17 Aug 2012 20:31:12 +0200 | |
changeset 48853 | ec82c33c75f8 |
parent 48611 | b34ff75c23a7 |
permissions | -rw-r--r-- |
16417 | 1 |
theory Relations imports Main begin |
10294 | 2 |
|
3 |
(*Id is only used in UNITY*) |
|
4 |
(*refl, antisym,trans,univalent,\<dots> ho hum*) |
|
5 |
||
6 |
text{* |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
7 |
@{thm[display] Id_def[no_vars]} |
10294 | 8 |
\rulename{Id_def} |
9 |
*} |
|
10 |
||
11 |
text{* |
|
47508 | 12 |
@{thm[display] relcomp_unfold[no_vars]} |
13 |
\rulename{relcomp_unfold} |
|
10294 | 14 |
*} |
15 |
||
16 |
text{* |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
17 |
@{thm[display] R_O_Id[no_vars]} |
10294 | 18 |
\rulename{R_O_Id} |
19 |
*} |
|
20 |
||
21 |
text{* |
|
47508 | 22 |
@{thm[display] relcomp_mono[no_vars]} |
23 |
\rulename{relcomp_mono} |
|
10294 | 24 |
*} |
25 |
||
26 |
text{* |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
27 |
@{thm[display] converse_iff[no_vars]} |
10294 | 28 |
\rulename{converse_iff} |
29 |
*} |
|
30 |
||
31 |
text{* |
|
47508 | 32 |
@{thm[display] converse_relcomp[no_vars]} |
33 |
\rulename{converse_relcomp} |
|
10294 | 34 |
*} |
35 |
||
36 |
text{* |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
37 |
@{thm[display] Image_iff[no_vars]} |
10294 | 38 |
\rulename{Image_iff} |
39 |
*} |
|
40 |
||
41 |
text{* |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
42 |
@{thm[display] Image_UN[no_vars]} |
10294 | 43 |
\rulename{Image_UN} |
44 |
*} |
|
45 |
||
46 |
text{* |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
47 |
@{thm[display] Domain_iff[no_vars]} |
10294 | 48 |
\rulename{Domain_iff} |
49 |
*} |
|
50 |
||
51 |
text{* |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
52 |
@{thm[display] Range_iff[no_vars]} |
10294 | 53 |
\rulename{Range_iff} |
54 |
*} |
|
55 |
||
56 |
text{* |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
57 |
@{thm[display] relpow.simps[no_vars]} |
10294 | 58 |
\rulename{relpow.simps} |
59 |
||
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
60 |
@{thm[display] rtrancl_refl[no_vars]} |
10294 | 61 |
\rulename{rtrancl_refl} |
62 |
||
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
63 |
@{thm[display] r_into_rtrancl[no_vars]} |
10294 | 64 |
\rulename{r_into_rtrancl} |
65 |
||
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
66 |
@{thm[display] rtrancl_trans[no_vars]} |
10294 | 67 |
\rulename{rtrancl_trans} |
68 |
||
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
69 |
@{thm[display] rtrancl_induct[no_vars]} |
10294 | 70 |
\rulename{rtrancl_induct} |
71 |
||
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
72 |
@{thm[display] rtrancl_idemp[no_vars]} |
10294 | 73 |
\rulename{rtrancl_idemp} |
74 |
||
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
75 |
@{thm[display] r_into_trancl[no_vars]} |
10294 | 76 |
\rulename{r_into_trancl} |
77 |
||
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
78 |
@{thm[display] trancl_trans[no_vars]} |
10294 | 79 |
\rulename{trancl_trans} |
80 |
||
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
81 |
@{thm[display] trancl_into_rtrancl[no_vars]} |
10294 | 82 |
\rulename{trancl_into_rtrancl} |
83 |
||
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
84 |
@{thm[display] trancl_converse[no_vars]} |
10294 | 85 |
\rulename{trancl_converse} |
86 |
*} |
|
87 |
||
88 |
text{*Relations. transitive closure*} |
|
89 |
||
11080
22855d091249
various revisions in response to comments from Tobias
paulson
parents:
10864
diff
changeset
|
90 |
lemma rtrancl_converseD: "(x,y) \<in> (r\<inverse>)\<^sup>* \<Longrightarrow> (y,x) \<in> r\<^sup>*" |
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
91 |
apply (erule rtrancl_induct) |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
92 |
txt{* |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
93 |
@{subgoals[display,indent=0,margin=65]} |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
94 |
*}; |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
95 |
apply (rule rtrancl_refl) |
11080
22855d091249
various revisions in response to comments from Tobias
paulson
parents:
10864
diff
changeset
|
96 |
apply (blast intro: rtrancl_trans) |
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
97 |
done |
10294 | 98 |
|
99 |
||
11080
22855d091249
various revisions in response to comments from Tobias
paulson
parents:
10864
diff
changeset
|
100 |
lemma rtrancl_converseI: "(y,x) \<in> r\<^sup>* \<Longrightarrow> (x,y) \<in> (r\<inverse>)\<^sup>*" |
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
101 |
apply (erule rtrancl_induct) |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
102 |
apply (rule rtrancl_refl) |
11080
22855d091249
various revisions in response to comments from Tobias
paulson
parents:
10864
diff
changeset
|
103 |
apply (blast intro: rtrancl_trans) |
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
104 |
done |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
105 |
|
11080
22855d091249
various revisions in response to comments from Tobias
paulson
parents:
10864
diff
changeset
|
106 |
lemma rtrancl_converse: "(r\<inverse>)\<^sup>* = (r\<^sup>*)\<inverse>" |
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
107 |
by (auto intro: rtrancl_converseI dest: rtrancl_converseD) |
10294 | 108 |
|
11080
22855d091249
various revisions in response to comments from Tobias
paulson
parents:
10864
diff
changeset
|
109 |
lemma rtrancl_converse: "(r\<inverse>)\<^sup>* = (r\<^sup>*)\<inverse>" |
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
110 |
apply (intro equalityI subsetI) |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
111 |
txt{* |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
112 |
after intro rules |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
113 |
|
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
114 |
@{subgoals[display,indent=0,margin=65]} |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
115 |
*}; |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
116 |
apply clarify |
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 splitting |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
119 |
@{subgoals[display,indent=0,margin=65]} |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
120 |
*}; |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
121 |
oops |
10294 | 122 |
|
123 |
||
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
124 |
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
|
125 |
apply (rule subsetI) |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
126 |
txt{* |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
127 |
@{subgoals[display,indent=0,margin=65]} |
10294 | 128 |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
129 |
after subsetI |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
130 |
*}; |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
131 |
apply clarify |
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 |
subgoals after clarify |
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 |
by blast |
10294 | 138 |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
139 |
|
10294 | 140 |
|
141 |
||
142 |
text{*rejects*} |
|
143 |
||
144 |
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
|
145 |
apply (blast) |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
146 |
done |
10294 | 147 |
|
148 |
text{*Pow, Inter too little used*} |
|
149 |
||
150 |
lemma "(A \<subset> B) = (A \<subseteq> B \<and> A \<noteq> B)" |
|
26806 | 151 |
apply (simp add: psubset_eq) |
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
152 |
done |
10294 | 153 |
|
154 |
end |