10341
|
1 |
(* ID: $Id$ *)
|
10295
|
2 |
theory Basic = Main:
|
|
3 |
|
|
4 |
lemma conj_rule: "\<lbrakk> P; Q \<rbrakk> \<Longrightarrow> P \<and> (Q \<and> P)"
|
|
5 |
apply (rule conjI)
|
|
6 |
apply assumption
|
|
7 |
apply (rule conjI)
|
|
8 |
apply assumption
|
|
9 |
apply assumption
|
|
10 |
done
|
|
11 |
|
|
12 |
|
|
13 |
lemma disj_swap: "P | Q \<Longrightarrow> Q | P"
|
|
14 |
apply (erule disjE)
|
|
15 |
apply (rule disjI2)
|
|
16 |
apply assumption
|
|
17 |
apply (rule disjI1)
|
|
18 |
apply assumption
|
|
19 |
done
|
|
20 |
|
|
21 |
lemma conj_swap: "P \<and> Q \<Longrightarrow> Q \<and> P"
|
|
22 |
apply (rule conjI)
|
|
23 |
apply (drule conjunct2)
|
|
24 |
apply assumption
|
|
25 |
apply (drule conjunct1)
|
|
26 |
apply assumption
|
|
27 |
done
|
|
28 |
|
|
29 |
lemma imp_uncurry: "P \<longrightarrow> Q \<longrightarrow> R \<Longrightarrow> P \<and> Q \<longrightarrow> R"
|
|
30 |
apply (rule impI)
|
|
31 |
apply (erule conjE)
|
|
32 |
apply (drule mp)
|
|
33 |
apply assumption
|
|
34 |
apply (drule mp)
|
|
35 |
apply assumption
|
|
36 |
apply assumption
|
|
37 |
done
|
|
38 |
|
10843
|
39 |
text {*NEW
|
|
40 |
by eliminates uses of assumption and done
|
|
41 |
*}
|
|
42 |
|
|
43 |
lemma imp_uncurry: "P \<longrightarrow> Q \<longrightarrow> R \<Longrightarrow> P \<and> Q \<longrightarrow> R"
|
|
44 |
apply (rule impI)
|
|
45 |
apply (erule conjE)
|
|
46 |
apply (drule mp)
|
|
47 |
apply assumption
|
|
48 |
by (drule mp)
|
|
49 |
|
|
50 |
|
10295
|
51 |
text {*
|
|
52 |
substitution
|
|
53 |
|
|
54 |
@{thm[display] ssubst}
|
|
55 |
\rulename{ssubst}
|
|
56 |
*};
|
|
57 |
|
|
58 |
lemma "\<lbrakk> x = f x; P(f x) \<rbrakk> \<Longrightarrow> P x"
|
10843
|
59 |
by (erule ssubst)
|
10295
|
60 |
|
|
61 |
text {*
|
|
62 |
also provable by simp (re-orients)
|
|
63 |
*};
|
|
64 |
|
|
65 |
lemma "\<lbrakk> x = f x; P (f x) (f x) x \<rbrakk> \<Longrightarrow> P x x x"
|
|
66 |
apply (erule ssubst)
|
|
67 |
back
|
|
68 |
back
|
|
69 |
back
|
|
70 |
back
|
|
71 |
apply assumption
|
|
72 |
done
|
|
73 |
|
|
74 |
text {*
|
|
75 |
proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ \isadigit{1}\isanewline
|
|
76 |
\isanewline
|
|
77 |
goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline
|
|
78 |
{\isasymlbrakk}x\ {\isacharequal}\ f\ x{\isacharsemicolon}\ P\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharparenleft}f\ x{\isacharparenright}\ x{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ x\ x\ x\isanewline
|
|
79 |
\ \isadigit{1}{\isachardot}\ P\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharparenleft}f\ x{\isacharparenright}\ x\ {\isasymLongrightarrow}\ P\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharparenleft}f\ x{\isacharparenright}
|
|
80 |
|
|
81 |
proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ \isadigit{1}\isanewline
|
|
82 |
\isanewline
|
|
83 |
goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline
|
|
84 |
{\isasymlbrakk}x\ {\isacharequal}\ f\ x{\isacharsemicolon}\ P\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharparenleft}f\ x{\isacharparenright}\ x{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ x\ x\ x\isanewline
|
|
85 |
\ \isadigit{1}{\isachardot}\ P\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharparenleft}f\ x{\isacharparenright}\ x\ {\isasymLongrightarrow}\ P\ x\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharparenleft}f\ x{\isacharparenright}
|
|
86 |
|
|
87 |
proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ \isadigit{1}\isanewline
|
|
88 |
\isanewline
|
|
89 |
goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline
|
|
90 |
{\isasymlbrakk}x\ {\isacharequal}\ f\ x{\isacharsemicolon}\ P\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharparenleft}f\ x{\isacharparenright}\ x{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ x\ x\ x\isanewline
|
|
91 |
\ \isadigit{1}{\isachardot}\ P\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharparenleft}f\ x{\isacharparenright}\ x\ {\isasymLongrightarrow}\ P\ {\isacharparenleft}f\ x{\isacharparenright}\ x\ {\isacharparenleft}f\ x{\isacharparenright}
|
|
92 |
|
|
93 |
proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ \isadigit{1}\isanewline
|
|
94 |
\isanewline
|
|
95 |
goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline
|
|
96 |
{\isasymlbrakk}x\ {\isacharequal}\ f\ x{\isacharsemicolon}\ P\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharparenleft}f\ x{\isacharparenright}\ x{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ x\ x\ x\isanewline
|
|
97 |
\ \isadigit{1}{\isachardot}\ P\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharparenleft}f\ x{\isacharparenright}\ x\ {\isasymLongrightarrow}\ P\ x\ x\ {\isacharparenleft}f\ x{\isacharparenright}
|
|
98 |
|
|
99 |
proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ \isadigit{1}\isanewline
|
|
100 |
\isanewline
|
|
101 |
goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline
|
|
102 |
{\isasymlbrakk}x\ {\isacharequal}\ f\ x{\isacharsemicolon}\ P\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharparenleft}f\ x{\isacharparenright}\ x{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ x\ x\ x\isanewline
|
|
103 |
\ \isadigit{1}{\isachardot}\ P\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharparenleft}f\ x{\isacharparenright}\ x\ {\isasymLongrightarrow}\ P\ {\isacharparenleft}f\ x{\isacharparenright}\ {\isacharparenleft}f\ x{\isacharparenright}\ x
|
|
104 |
*};
|
|
105 |
|
|
106 |
lemma "\<lbrakk> x = f x; P (f x) (f x) x \<rbrakk> \<Longrightarrow> P x x x"
|
|
107 |
apply (erule ssubst, assumption)
|
|
108 |
done
|
|
109 |
|
10843
|
110 |
text{*
|
|
111 |
or better still NEW
|
|
112 |
*}
|
|
113 |
|
10295
|
114 |
lemma "\<lbrakk> x = f x; P (f x) (f x) x \<rbrakk> \<Longrightarrow> P x x x"
|
10843
|
115 |
by (erule ssubst)
|
|
116 |
|
|
117 |
|
|
118 |
text{*NEW*}
|
|
119 |
lemma "\<lbrakk> x = f x; P (f x) (f x) x \<rbrakk> \<Longrightarrow> P x x x"
|
|
120 |
apply (erule_tac P="\<lambda>u. P u u x" in ssubst)
|
|
121 |
apply (assumption)
|
10295
|
122 |
done
|
|
123 |
|
|
124 |
|
10843
|
125 |
lemma "\<lbrakk> x = f x; P (f x) (f x) x \<rbrakk> \<Longrightarrow> P x x x"
|
|
126 |
by (erule_tac P="\<lambda>u. P u u x" in ssubst)
|
|
127 |
|
|
128 |
|
10295
|
129 |
text {*
|
|
130 |
negation
|
|
131 |
|
|
132 |
@{thm[display] notI}
|
|
133 |
\rulename{notI}
|
|
134 |
|
|
135 |
@{thm[display] notE}
|
|
136 |
\rulename{notE}
|
|
137 |
|
|
138 |
@{thm[display] classical}
|
|
139 |
\rulename{classical}
|
|
140 |
|
|
141 |
@{thm[display] contrapos_pp}
|
|
142 |
\rulename{contrapos_pp}
|
|
143 |
|
|
144 |
@{thm[display] contrapos_np}
|
|
145 |
\rulename{contrapos_np}
|
|
146 |
|
|
147 |
@{thm[display] contrapos_nn}
|
|
148 |
\rulename{contrapos_nn}
|
|
149 |
*};
|
|
150 |
|
|
151 |
|
|
152 |
lemma "\<lbrakk>\<not>(P\<longrightarrow>Q); \<not>(R\<longrightarrow>Q)\<rbrakk> \<Longrightarrow> R"
|
|
153 |
apply (erule_tac Q="R\<longrightarrow>Q" in contrapos_np)
|
|
154 |
txt{*
|
|
155 |
proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ {\isadigit{1}}\isanewline
|
|
156 |
\isanewline
|
|
157 |
goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline
|
|
158 |
{\isasymlbrakk}{\isasymnot}\ {\isacharparenleft}P\ {\isasymlongrightarrow}\ Q{\isacharparenright}{\isacharsemicolon}\ {\isasymnot}\ {\isacharparenleft}R\ {\isasymlongrightarrow}\ Q{\isacharparenright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ R\isanewline
|
|
159 |
\ {\isadigit{1}}{\isachardot}\ {\isasymlbrakk}{\isasymnot}\ {\isacharparenleft}P\ {\isasymlongrightarrow}\ Q{\isacharparenright}{\isacharsemicolon}\ {\isasymnot}\ R{\isasymrbrakk}\ {\isasymLongrightarrow}\ R\ {\isasymlongrightarrow}\ Q
|
|
160 |
*}
|
|
161 |
|
|
162 |
apply intro
|
|
163 |
txt{*
|
|
164 |
proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ {\isadigit{3}}\isanewline
|
|
165 |
\isanewline
|
|
166 |
goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline
|
|
167 |
{\isasymlbrakk}{\isasymnot}\ {\isacharparenleft}P\ {\isasymlongrightarrow}\ Q{\isacharparenright}{\isacharsemicolon}\ {\isasymnot}\ {\isacharparenleft}R\ {\isasymlongrightarrow}\ Q{\isacharparenright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ R\isanewline
|
|
168 |
\ {\isadigit{1}}{\isachardot}\ {\isasymlbrakk}{\isasymnot}\ {\isacharparenleft}P\ {\isasymlongrightarrow}\ Q{\isacharparenright}{\isacharsemicolon}\ {\isasymnot}\ R{\isacharsemicolon}\ R{\isasymrbrakk}\ {\isasymLongrightarrow}\ Q
|
|
169 |
*}
|
10843
|
170 |
by (erule notE)
|
|
171 |
text{*NEW*}
|
|
172 |
|
10295
|
173 |
|
|
174 |
|
|
175 |
lemma "(P \<or> Q) \<and> R \<Longrightarrow> P \<or> Q \<and> R"
|
|
176 |
apply intro
|
|
177 |
txt{*
|
|
178 |
proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ {\isadigit{1}}\isanewline
|
|
179 |
\isanewline
|
|
180 |
goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline
|
|
181 |
{\isacharparenleft}P\ {\isasymor}\ Q{\isacharparenright}\ {\isasymand}\ R\ {\isasymLongrightarrow}\ P\ {\isasymor}\ Q\ {\isasymand}\ R\isanewline
|
|
182 |
\ {\isadigit{1}}{\isachardot}\ {\isasymlbrakk}{\isacharparenleft}P\ {\isasymor}\ Q{\isacharparenright}\ {\isasymand}\ R{\isacharsemicolon}\ {\isasymnot}\ {\isacharparenleft}Q\ {\isasymand}\ R{\isacharparenright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ P
|
|
183 |
*}
|
|
184 |
|
|
185 |
apply (elim conjE disjE)
|
|
186 |
apply assumption
|
|
187 |
|
|
188 |
txt{*
|
|
189 |
proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ {\isadigit{4}}\isanewline
|
|
190 |
\isanewline
|
|
191 |
goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline
|
|
192 |
{\isacharparenleft}P\ {\isasymor}\ Q{\isacharparenright}\ {\isasymand}\ R\ {\isasymLongrightarrow}\ P\ {\isasymor}\ Q\ {\isasymand}\ R\isanewline
|
|
193 |
\ {\isadigit{1}}{\isachardot}\ {\isasymlbrakk}{\isasymnot}\ {\isacharparenleft}Q\ {\isasymand}\ R{\isacharparenright}{\isacharsemicolon}\ R{\isacharsemicolon}\ Q{\isasymrbrakk}\ {\isasymLongrightarrow}\ P
|
|
194 |
*}
|
|
195 |
|
10843
|
196 |
by (erule contrapos_np, rule conjI)
|
|
197 |
text{*NEW
|
10295
|
198 |
proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ {\isadigit{6}}\isanewline
|
|
199 |
\isanewline
|
|
200 |
goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline
|
|
201 |
{\isacharparenleft}P\ {\isasymor}\ Q{\isacharparenright}\ {\isasymand}\ R\ {\isasymLongrightarrow}\ P\ {\isasymor}\ Q\ {\isasymand}\ R\isanewline
|
|
202 |
\ {\isadigit{1}}{\isachardot}\ {\isasymlbrakk}R{\isacharsemicolon}\ Q{\isacharsemicolon}\ {\isasymnot}\ P{\isasymrbrakk}\ {\isasymLongrightarrow}\ Q\isanewline
|
|
203 |
\ {\isadigit{2}}{\isachardot}\ {\isasymlbrakk}R{\isacharsemicolon}\ Q{\isacharsemicolon}\ {\isasymnot}\ P{\isasymrbrakk}\ {\isasymLongrightarrow}\ R
|
|
204 |
*}
|
|
205 |
|
|
206 |
|
|
207 |
text{*Quantifiers*}
|
|
208 |
|
|
209 |
lemma "\<forall>x. P x \<longrightarrow> P x"
|
|
210 |
apply (rule allI)
|
10843
|
211 |
by (rule impI)
|
|
212 |
text{*NEW*}
|
10295
|
213 |
|
|
214 |
lemma "(\<forall>x. P \<longrightarrow> Q x) \<Longrightarrow> P \<longrightarrow> (\<forall>x. Q x)"
|
10843
|
215 |
apply (rule impI, rule allI)
|
10295
|
216 |
apply (drule spec)
|
10843
|
217 |
by (drule mp)
|
|
218 |
text{*NEW*}
|
10295
|
219 |
|
10843
|
220 |
lemma "\<lbrakk>\<forall>x. P x \<longrightarrow> P (h x); P a\<rbrakk> \<Longrightarrow> P(h (h a))"
|
10295
|
221 |
apply (frule spec)
|
|
222 |
apply (drule mp, assumption)
|
|
223 |
apply (drule spec)
|
10843
|
224 |
by (drule mp)
|
|
225 |
text{*NEW*}
|
|
226 |
|
10295
|
227 |
|
|
228 |
text
|
|
229 |
{*
|
|
230 |
proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ {\isadigit{1}}\isanewline
|
|
231 |
\isanewline
|
|
232 |
goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline
|
|
233 |
{\isasymlbrakk}{\isasymforall}x{\isachardot}\ P\ x\ {\isasymlongrightarrow}\ P\ {\isacharparenleft}f\ x{\isacharparenright}{\isacharsemicolon}\ P\ a{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ {\isacharparenleft}f\ {\isacharparenleft}f\ a{\isacharparenright}{\isacharparenright}\isanewline
|
|
234 |
\ {\isadigit{1}}{\isachardot}\ {\isasymlbrakk}{\isasymforall}x{\isachardot}\ P\ x\ {\isasymlongrightarrow}\ P\ {\isacharparenleft}f\ x{\isacharparenright}{\isacharsemicolon}\ P\ a{\isacharsemicolon}\ P\ {\isacharquery}x\ {\isasymlongrightarrow}\ P\ {\isacharparenleft}f\ {\isacharquery}x{\isacharparenright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ {\isacharparenleft}f\ {\isacharparenleft}f\ a{\isacharparenright}{\isacharparenright}
|
|
235 |
*}
|
|
236 |
|
|
237 |
text{*
|
|
238 |
proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ {\isadigit{3}}\isanewline
|
|
239 |
\isanewline
|
|
240 |
goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline
|
|
241 |
{\isasymlbrakk}{\isasymforall}x{\isachardot}\ P\ x\ {\isasymlongrightarrow}\ P\ {\isacharparenleft}f\ x{\isacharparenright}{\isacharsemicolon}\ P\ a{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ {\isacharparenleft}f\ {\isacharparenleft}f\ a{\isacharparenright}{\isacharparenright}\isanewline
|
|
242 |
\ {\isadigit{1}}{\isachardot}\ {\isasymlbrakk}{\isasymforall}x{\isachardot}\ P\ x\ {\isasymlongrightarrow}\ P\ {\isacharparenleft}f\ x{\isacharparenright}{\isacharsemicolon}\ P\ a{\isacharsemicolon}\ P\ {\isacharparenleft}f\ a{\isacharparenright}{\isasymrbrakk}\ {\isasymLongrightarrow}\ P\ {\isacharparenleft}f\ {\isacharparenleft}f\ a{\isacharparenright}{\isacharparenright}
|
|
243 |
*}
|
|
244 |
|
|
245 |
lemma "\<lbrakk>\<forall>x. P x \<longrightarrow> P (f x); P a\<rbrakk> \<Longrightarrow> P(f (f a))"
|
|
246 |
by blast
|
|
247 |
|
10843
|
248 |
text{*NEW
|
|
249 |
Hilbert-epsilon theorems*}
|
|
250 |
|
|
251 |
text{*
|
|
252 |
@{thm[display] some_equality[no_vars]}
|
|
253 |
\rulename{some_equality}
|
|
254 |
|
|
255 |
@{thm[display] someI[no_vars]}
|
|
256 |
\rulename{someI}
|
|
257 |
|
|
258 |
@{thm[display] someI2[no_vars]}
|
|
259 |
\rulename{someI2}
|
|
260 |
|
|
261 |
needed for examples
|
|
262 |
|
|
263 |
@{thm[display] inv_def[no_vars]}
|
|
264 |
\rulename{inv_def}
|
|
265 |
|
|
266 |
@{thm[display] Least_def[no_vars]}
|
|
267 |
\rulename{Least_def}
|
|
268 |
|
|
269 |
@{thm[display] order_antisym[no_vars]}
|
|
270 |
\rulename{order_antisym}
|
|
271 |
*}
|
|
272 |
|
|
273 |
|
|
274 |
lemma "inv Suc (Suc x) = x"
|
|
275 |
by (simp add: inv_def)
|
|
276 |
|
|
277 |
text{*but we know nothing about inv Suc 0*}
|
|
278 |
|
|
279 |
theorem Least_equality:
|
|
280 |
"\<lbrakk> P (k::nat); \<forall>x. P x \<longrightarrow> k \<le> x \<rbrakk> \<Longrightarrow> (LEAST x. P(x)) = k"
|
|
281 |
apply (simp add: Least_def)
|
|
282 |
|
|
283 |
txt{*omit maybe?
|
|
284 |
@{subgoals[display,indent=0,margin=65]}
|
|
285 |
*};
|
|
286 |
|
|
287 |
apply (rule some_equality)
|
|
288 |
|
|
289 |
txt{*
|
|
290 |
@{subgoals[display,indent=0,margin=65]}
|
|
291 |
|
|
292 |
first subgoal is existence; second is uniqueness
|
|
293 |
*};
|
|
294 |
by (auto intro: order_antisym)
|
|
295 |
|
|
296 |
|
|
297 |
theorem axiom_of_choice:
|
|
298 |
"(\<forall>x. \<exists> y. P x y) \<Longrightarrow> \<exists>f. \<forall>x. P x (f x)"
|
|
299 |
apply (rule exI, rule allI)
|
|
300 |
|
|
301 |
txt{*
|
|
302 |
@{subgoals[display,indent=0,margin=65]}
|
|
303 |
|
|
304 |
state after intro rules
|
|
305 |
*};
|
|
306 |
apply (drule spec, erule exE)
|
|
307 |
|
|
308 |
txt{*
|
|
309 |
@{subgoals[display,indent=0,margin=65]}
|
|
310 |
|
|
311 |
applying @text{someI} automatically instantiates
|
|
312 |
@{term f} to @{term "\<lambda>x. SOME y. P x y"}
|
|
313 |
*};
|
|
314 |
|
|
315 |
by (rule someI)
|
|
316 |
|
|
317 |
(*both can be done by blast, which however hasn't been introduced yet*)
|
|
318 |
lemma "[| P (k::nat); \<forall>x. P x \<longrightarrow> k \<le> x |] ==> (LEAST x. P(x)) = k";
|
|
319 |
apply (simp add: Least_def)
|
|
320 |
by (blast intro: some_equality order_antisym);
|
|
321 |
|
|
322 |
theorem axiom_of_choice: "(\<forall>x. \<exists> y. P x y) \<Longrightarrow> \<exists>f. \<forall>x. P x (f x)"
|
|
323 |
apply (rule exI [of _ "\<lambda>x. SOME y. P x y"])
|
|
324 |
by (blast intro: someI);
|
|
325 |
|
|
326 |
text{*end of NEW material*}
|
|
327 |
|
10295
|
328 |
lemma "(\<exists>x. P x) \<or> (\<exists>x. Q x) \<Longrightarrow> \<exists>x. P x \<or> Q x"
|
|
329 |
apply elim
|
|
330 |
apply intro
|
|
331 |
apply assumption
|
|
332 |
apply (intro exI disjI2) (*or else we get disjCI*)
|
|
333 |
apply assumption
|
|
334 |
done
|
|
335 |
|
|
336 |
lemma "(P\<longrightarrow>Q) \<or> (Q\<longrightarrow>P)"
|
|
337 |
apply intro
|
|
338 |
apply elim
|
|
339 |
apply assumption
|
|
340 |
done
|
|
341 |
|
|
342 |
lemma "(P\<or>Q)\<and>(P\<or>R) \<Longrightarrow> P \<or> (Q\<and>R)"
|
|
343 |
apply intro
|
|
344 |
apply (elim conjE disjE)
|
|
345 |
apply blast
|
|
346 |
apply blast
|
|
347 |
apply blast
|
|
348 |
apply blast
|
|
349 |
(*apply elim*)
|
|
350 |
done
|
|
351 |
|
|
352 |
lemma "(\<exists>x. P \<and> Q x) \<Longrightarrow> P \<and> (\<exists>x. Q x)"
|
|
353 |
apply (erule exE)
|
|
354 |
apply (erule conjE)
|
|
355 |
apply (rule conjI)
|
|
356 |
apply assumption
|
|
357 |
apply (rule exI)
|
|
358 |
apply assumption
|
|
359 |
done
|
|
360 |
|
|
361 |
lemma "(\<exists>x. P x) \<and> (\<exists>x. Q x) \<Longrightarrow> \<exists>x. P x \<and> Q x"
|
|
362 |
apply (erule conjE)
|
|
363 |
apply (erule exE)
|
|
364 |
apply (erule exE)
|
|
365 |
apply (rule exI)
|
|
366 |
apply (rule conjI)
|
|
367 |
apply assumption
|
|
368 |
oops
|
|
369 |
|
|
370 |
lemma "\<forall> z. R z z \<Longrightarrow> \<exists>x. \<forall> y. R x y"
|
|
371 |
apply (rule exI)
|
|
372 |
apply (rule allI)
|
|
373 |
apply (drule spec)
|
|
374 |
oops
|
|
375 |
|
|
376 |
lemma "\<forall>x. \<exists> y. x=y"
|
|
377 |
apply (rule allI)
|
|
378 |
apply (rule exI)
|
|
379 |
apply (rule refl)
|
|
380 |
done
|
|
381 |
|
|
382 |
lemma "\<exists>x. \<forall> y. x=y"
|
|
383 |
apply (rule exI)
|
|
384 |
apply (rule allI)
|
|
385 |
oops
|
|
386 |
|
|
387 |
end
|