1(*  Title:      FOL/IFOL.thy
2    Author:     Lawrence C Paulson and Markus Wenzel
3*)
4
5section \<open>Intuitionistic first-order logic\<close>
6
7theory IFOL
8imports Pure
9begin
10
11ML_file \<open>~~/src/Tools/misc_legacy.ML\<close>
12ML_file \<open>~~/src/Provers/splitter.ML\<close>
13ML_file \<open>~~/src/Provers/hypsubst.ML\<close>
14ML_file \<open>~~/src/Tools/IsaPlanner/zipper.ML\<close>
15ML_file \<open>~~/src/Tools/IsaPlanner/isand.ML\<close>
16ML_file \<open>~~/src/Tools/IsaPlanner/rw_inst.ML\<close>
17ML_file \<open>~~/src/Provers/quantifier1.ML\<close>
18ML_file \<open>~~/src/Tools/intuitionistic.ML\<close>
19ML_file \<open>~~/src/Tools/project_rule.ML\<close>
20ML_file \<open>~~/src/Tools/atomize_elim.ML\<close>
21
22
23subsection \<open>Syntax and axiomatic basis\<close>
24
25setup Pure_Thy.old_appl_syntax_setup
26setup \<open>Proofterm.set_preproc (Proof_Rewrite_Rules.standard_preproc [])\<close>
27
28class "term"
29default_sort \<open>term\<close>
30
31typedecl o
32
33judgment
34  Trueprop :: \<open>o \<Rightarrow> prop\<close>  (\<open>(_)\<close> 5)
35
36
37subsubsection \<open>Equality\<close>
38
39axiomatization
40  eq :: \<open>['a, 'a] \<Rightarrow> o\<close>  (infixl \<open>=\<close> 50)
41where
42  refl: \<open>a = a\<close> and
43  subst: \<open>a = b \<Longrightarrow> P(a) \<Longrightarrow> P(b)\<close>
44
45
46subsubsection \<open>Propositional logic\<close>
47
48axiomatization
49  False :: \<open>o\<close> and
50  conj :: \<open>[o, o] => o\<close>  (infixr \<open>\<and>\<close> 35) and
51  disj :: \<open>[o, o] => o\<close>  (infixr \<open>\<or>\<close> 30) and
52  imp :: \<open>[o, o] => o\<close>  (infixr \<open>\<longrightarrow>\<close> 25)
53where
54  conjI: \<open>\<lbrakk>P;  Q\<rbrakk> \<Longrightarrow> P \<and> Q\<close> and
55  conjunct1: \<open>P \<and> Q \<Longrightarrow> P\<close> and
56  conjunct2: \<open>P \<and> Q \<Longrightarrow> Q\<close> and
57
58  disjI1: \<open>P \<Longrightarrow> P \<or> Q\<close> and
59  disjI2: \<open>Q \<Longrightarrow> P \<or> Q\<close> and
60  disjE: \<open>\<lbrakk>P \<or> Q; P \<Longrightarrow> R; Q \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R\<close> and
61
62  impI: \<open>(P \<Longrightarrow> Q) \<Longrightarrow> P \<longrightarrow> Q\<close> and
63  mp: \<open>\<lbrakk>P \<longrightarrow> Q; P\<rbrakk> \<Longrightarrow> Q\<close> and
64
65  FalseE: \<open>False \<Longrightarrow> P\<close>
66
67
68subsubsection \<open>Quantifiers\<close>
69
70axiomatization
71  All :: \<open>('a \<Rightarrow> o) \<Rightarrow> o\<close>  (binder \<open>\<forall>\<close> 10) and
72  Ex :: \<open>('a \<Rightarrow> o) \<Rightarrow> o\<close>  (binder \<open>\<exists>\<close> 10)
73where
74  allI: \<open>(\<And>x. P(x)) \<Longrightarrow> (\<forall>x. P(x))\<close> and
75  spec: \<open>(\<forall>x. P(x)) \<Longrightarrow> P(x)\<close> and
76  exI: \<open>P(x) \<Longrightarrow> (\<exists>x. P(x))\<close> and
77  exE: \<open>\<lbrakk>\<exists>x. P(x); \<And>x. P(x) \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R\<close>
78
79
80subsubsection \<open>Definitions\<close>
81
82definition \<open>True \<equiv> False \<longrightarrow> False\<close>
83
84definition Not (\<open>\<not> _\<close> [40] 40)
85  where not_def: \<open>\<not> P \<equiv> P \<longrightarrow> False\<close>
86
87definition iff  (infixr \<open>\<longleftrightarrow>\<close> 25)
88  where \<open>P \<longleftrightarrow> Q \<equiv> (P \<longrightarrow> Q) \<and> (Q \<longrightarrow> P)\<close>
89
90definition Ex1 :: \<open>('a \<Rightarrow> o) \<Rightarrow> o\<close>  (binder \<open>\<exists>!\<close> 10)
91  where ex1_def: \<open>\<exists>!x. P(x) \<equiv> \<exists>x. P(x) \<and> (\<forall>y. P(y) \<longrightarrow> y = x)\<close>
92
93axiomatization where  \<comment> \<open>Reflection, admissible\<close>
94  eq_reflection: \<open>(x = y) \<Longrightarrow> (x \<equiv> y)\<close> and
95  iff_reflection: \<open>(P \<longleftrightarrow> Q) \<Longrightarrow> (P \<equiv> Q)\<close>
96
97abbreviation not_equal :: \<open>['a, 'a] \<Rightarrow> o\<close>  (infixl \<open>\<noteq>\<close> 50)
98  where \<open>x \<noteq> y \<equiv> \<not> (x = y)\<close>
99
100
101subsubsection \<open>Old-style ASCII syntax\<close>
102
103notation (ASCII)
104  not_equal  (infixl \<open>~=\<close> 50) and
105  Not  (\<open>~ _\<close> [40] 40) and
106  conj  (infixr \<open>&\<close> 35) and
107  disj  (infixr \<open>|\<close> 30) and
108  All  (binder \<open>ALL \<close> 10) and
109  Ex  (binder \<open>EX \<close> 10) and
110  Ex1  (binder \<open>EX! \<close> 10) and
111  imp  (infixr \<open>-->\<close> 25) and
112  iff  (infixr \<open><->\<close> 25)
113
114
115subsection \<open>Lemmas and proof tools\<close>
116
117lemmas strip = impI allI
118
119lemma TrueI: \<open>True\<close>
120  unfolding True_def by (rule impI)
121
122
123subsubsection \<open>Sequent-style elimination rules for \<open>\<and>\<close> \<open>\<longrightarrow>\<close> and \<open>\<forall>\<close>\<close>
124
125lemma conjE:
126  assumes major: \<open>P \<and> Q\<close>
127    and r: \<open>\<lbrakk>P; Q\<rbrakk> \<Longrightarrow> R\<close>
128  shows \<open>R\<close>
129  apply (rule r)
130   apply (rule major [THEN conjunct1])
131  apply (rule major [THEN conjunct2])
132  done
133
134lemma impE:
135  assumes major: \<open>P \<longrightarrow> Q\<close>
136    and \<open>P\<close>
137  and r: \<open>Q \<Longrightarrow> R\<close>
138  shows \<open>R\<close>
139  apply (rule r)
140  apply (rule major [THEN mp])
141  apply (rule \<open>P\<close>)
142  done
143
144lemma allE:
145  assumes major: \<open>\<forall>x. P(x)\<close>
146    and r: \<open>P(x) \<Longrightarrow> R\<close>
147  shows \<open>R\<close>
148  apply (rule r)
149  apply (rule major [THEN spec])
150  done
151
152text \<open>Duplicates the quantifier; for use with \<^ML>\<open>eresolve_tac\<close>.\<close>
153lemma all_dupE:
154  assumes major: \<open>\<forall>x. P(x)\<close>
155    and r: \<open>\<lbrakk>P(x); \<forall>x. P(x)\<rbrakk> \<Longrightarrow> R\<close>
156  shows \<open>R\<close>
157  apply (rule r)
158   apply (rule major [THEN spec])
159  apply (rule major)
160  done
161
162
163subsubsection \<open>Negation rules, which translate between \<open>\<not> P\<close> and \<open>P \<longrightarrow> False\<close>\<close>
164
165lemma notI: \<open>(P \<Longrightarrow> False) \<Longrightarrow> \<not> P\<close>
166  unfolding not_def by (erule impI)
167
168lemma notE: \<open>\<lbrakk>\<not> P; P\<rbrakk> \<Longrightarrow> R\<close>
169  unfolding not_def by (erule mp [THEN FalseE])
170
171lemma rev_notE: \<open>\<lbrakk>P; \<not> P\<rbrakk> \<Longrightarrow> R\<close>
172  by (erule notE)
173
174text \<open>This is useful with the special implication rules for each kind of \<open>P\<close>.\<close>
175lemma not_to_imp:
176  assumes \<open>\<not> P\<close>
177    and r: \<open>P \<longrightarrow> False \<Longrightarrow> Q\<close>
178  shows \<open>Q\<close>
179  apply (rule r)
180  apply (rule impI)
181  apply (erule notE [OF \<open>\<not> P\<close>])
182  done
183
184text \<open>
185  For substitution into an assumption \<open>P\<close>, reduce \<open>Q\<close> to \<open>P \<longrightarrow> Q\<close>, substitute into this implication, then apply \<open>impI\<close> to
186  move \<open>P\<close> back into the assumptions.
187\<close>
188lemma rev_mp: \<open>\<lbrakk>P; P \<longrightarrow> Q\<rbrakk> \<Longrightarrow> Q\<close>
189  by (erule mp)
190
191text \<open>Contrapositive of an inference rule.\<close>
192lemma contrapos:
193  assumes major: \<open>\<not> Q\<close>
194    and minor: \<open>P \<Longrightarrow> Q\<close>
195  shows \<open>\<not> P\<close>
196  apply (rule major [THEN notE, THEN notI])
197  apply (erule minor)
198  done
199
200
201subsubsection \<open>Modus Ponens Tactics\<close>
202
203text \<open>
204  Finds \<open>P \<longrightarrow> Q\<close> and P in the assumptions, replaces implication by
205  \<open>Q\<close>.
206\<close>
207ML \<open>
208  fun mp_tac ctxt i =
209    eresolve_tac ctxt @{thms notE impE} i THEN assume_tac ctxt i;
210  fun eq_mp_tac ctxt i =
211    eresolve_tac ctxt @{thms notE impE} i THEN eq_assume_tac i;
212\<close>
213
214
215subsection \<open>If-and-only-if\<close>
216
217lemma iffI: \<open>\<lbrakk>P \<Longrightarrow> Q; Q \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P \<longleftrightarrow> Q\<close>
218  apply (unfold iff_def)
219  apply (rule conjI)
220   apply (erule impI)
221  apply (erule impI)
222  done
223
224lemma iffE:
225  assumes major: \<open>P \<longleftrightarrow> Q\<close>
226    and r: \<open>P \<longrightarrow> Q \<Longrightarrow> Q \<longrightarrow> P \<Longrightarrow> R\<close>
227  shows \<open>R\<close>
228  apply (insert major, unfold iff_def)
229  apply (erule conjE)
230  apply (erule r)
231  apply assumption
232  done
233
234
235subsubsection \<open>Destruct rules for \<open>\<longleftrightarrow>\<close> similar to Modus Ponens\<close>
236
237lemma iffD1: \<open>\<lbrakk>P \<longleftrightarrow> Q; P\<rbrakk> \<Longrightarrow> Q\<close>
238  apply (unfold iff_def)
239  apply (erule conjunct1 [THEN mp])
240  apply assumption
241  done
242
243lemma iffD2: \<open>\<lbrakk>P \<longleftrightarrow> Q; Q\<rbrakk> \<Longrightarrow> P\<close>
244  apply (unfold iff_def)
245  apply (erule conjunct2 [THEN mp])
246  apply assumption
247  done
248
249lemma rev_iffD1: \<open>\<lbrakk>P; P \<longleftrightarrow> Q\<rbrakk> \<Longrightarrow> Q\<close>
250  apply (erule iffD1)
251  apply assumption
252  done
253
254lemma rev_iffD2: \<open>\<lbrakk>Q; P \<longleftrightarrow> Q\<rbrakk> \<Longrightarrow> P\<close>
255  apply (erule iffD2)
256  apply assumption
257  done
258
259lemma iff_refl: \<open>P \<longleftrightarrow> P\<close>
260  by (rule iffI)
261
262lemma iff_sym: \<open>Q \<longleftrightarrow> P \<Longrightarrow> P \<longleftrightarrow> Q\<close>
263  apply (erule iffE)
264  apply (rule iffI)
265  apply (assumption | erule mp)+
266  done
267
268lemma iff_trans: \<open>\<lbrakk>P \<longleftrightarrow> Q; Q \<longleftrightarrow> R\<rbrakk> \<Longrightarrow> P \<longleftrightarrow> R\<close>
269  apply (rule iffI)
270  apply (assumption | erule iffE | erule (1) notE impE)+
271  done
272
273
274subsection \<open>Unique existence\<close>
275
276text \<open>
277  NOTE THAT the following 2 quantifications:
278
279    \<^item> \<open>\<exists>!x\<close> such that [\<open>\<exists>!y\<close> such that P(x,y)]   (sequential)
280    \<^item> \<open>\<exists>!x,y\<close> such that P(x,y)                   (simultaneous)
281
282  do NOT mean the same thing. The parser treats \<open>\<exists>!x y.P(x,y)\<close> as sequential.
283\<close>
284
285lemma ex1I: \<open>P(a) \<Longrightarrow> (\<And>x. P(x) \<Longrightarrow> x = a) \<Longrightarrow> \<exists>!x. P(x)\<close>
286  apply (unfold ex1_def)
287  apply (assumption | rule exI conjI allI impI)+
288  done
289
290text \<open>Sometimes easier to use: the premises have no shared variables. Safe!\<close>
291lemma ex_ex1I: \<open>\<exists>x. P(x) \<Longrightarrow> (\<And>x y. \<lbrakk>P(x); P(y)\<rbrakk> \<Longrightarrow> x = y) \<Longrightarrow> \<exists>!x. P(x)\<close>
292  apply (erule exE)
293  apply (rule ex1I)
294   apply assumption
295  apply assumption
296  done
297
298lemma ex1E: \<open>\<exists>! x. P(x) \<Longrightarrow> (\<And>x. \<lbrakk>P(x); \<forall>y. P(y) \<longrightarrow> y = x\<rbrakk> \<Longrightarrow> R) \<Longrightarrow> R\<close>
299  apply (unfold ex1_def)
300  apply (assumption | erule exE conjE)+
301  done
302
303
304subsubsection \<open>\<open>\<longleftrightarrow>\<close> congruence rules for simplification\<close>
305
306text \<open>Use \<open>iffE\<close> on a premise. For \<open>conj_cong\<close>, \<open>imp_cong\<close>, \<open>all_cong\<close>, \<open>ex_cong\<close>.\<close>
307ML \<open>
308  fun iff_tac ctxt prems i =
309    resolve_tac ctxt (prems RL @{thms iffE}) i THEN
310    REPEAT1 (eresolve_tac ctxt @{thms asm_rl mp} i);
311\<close>
312
313method_setup iff =
314  \<open>Attrib.thms >>
315    (fn prems => fn ctxt => SIMPLE_METHOD' (iff_tac ctxt prems))\<close>
316
317lemma conj_cong:
318  assumes \<open>P \<longleftrightarrow> P'\<close>
319    and \<open>P' \<Longrightarrow> Q \<longleftrightarrow> Q'\<close>
320  shows \<open>(P \<and> Q) \<longleftrightarrow> (P' \<and> Q')\<close>
321  apply (insert assms)
322  apply (assumption | rule iffI conjI | erule iffE conjE mp | iff assms)+
323  done
324
325text \<open>Reversed congruence rule!  Used in ZF/Order.\<close>
326lemma conj_cong2:
327  assumes \<open>P \<longleftrightarrow> P'\<close>
328    and \<open>P' \<Longrightarrow> Q \<longleftrightarrow> Q'\<close>
329  shows \<open>(Q \<and> P) \<longleftrightarrow> (Q' \<and> P')\<close>
330  apply (insert assms)
331  apply (assumption | rule iffI conjI | erule iffE conjE mp | iff assms)+
332  done
333
334lemma disj_cong:
335  assumes \<open>P \<longleftrightarrow> P'\<close> and \<open>Q \<longleftrightarrow> Q'\<close>
336  shows \<open>(P \<or> Q) \<longleftrightarrow> (P' \<or> Q')\<close>
337  apply (insert assms)
338  apply (erule iffE disjE disjI1 disjI2 |
339    assumption | rule iffI | erule (1) notE impE)+
340  done
341
342lemma imp_cong:
343  assumes \<open>P \<longleftrightarrow> P'\<close>
344    and \<open>P' \<Longrightarrow> Q \<longleftrightarrow> Q'\<close>
345  shows \<open>(P \<longrightarrow> Q) \<longleftrightarrow> (P' \<longrightarrow> Q')\<close>
346  apply (insert assms)
347  apply (assumption | rule iffI impI | erule iffE | erule (1) notE impE | iff assms)+
348  done
349
350lemma iff_cong: \<open>\<lbrakk>P \<longleftrightarrow> P'; Q \<longleftrightarrow> Q'\<rbrakk> \<Longrightarrow> (P \<longleftrightarrow> Q) \<longleftrightarrow> (P' \<longleftrightarrow> Q')\<close>
351  apply (erule iffE | assumption | rule iffI | erule (1) notE impE)+
352  done
353
354lemma not_cong: \<open>P \<longleftrightarrow> P' \<Longrightarrow> \<not> P \<longleftrightarrow> \<not> P'\<close>
355  apply (assumption | rule iffI notI | erule (1) notE impE | erule iffE notE)+
356  done
357
358lemma all_cong:
359  assumes \<open>\<And>x. P(x) \<longleftrightarrow> Q(x)\<close>
360  shows \<open>(\<forall>x. P(x)) \<longleftrightarrow> (\<forall>x. Q(x))\<close>
361  apply (assumption | rule iffI allI | erule (1) notE impE | erule allE | iff assms)+
362  done
363
364lemma ex_cong:
365  assumes \<open>\<And>x. P(x) \<longleftrightarrow> Q(x)\<close>
366  shows \<open>(\<exists>x. P(x)) \<longleftrightarrow> (\<exists>x. Q(x))\<close>
367  apply (erule exE | assumption | rule iffI exI | erule (1) notE impE | iff assms)+
368  done
369
370lemma ex1_cong:
371  assumes \<open>\<And>x. P(x) \<longleftrightarrow> Q(x)\<close>
372  shows \<open>(\<exists>!x. P(x)) \<longleftrightarrow> (\<exists>!x. Q(x))\<close>
373  apply (erule ex1E spec [THEN mp] | assumption | rule iffI ex1I | erule (1) notE impE | iff assms)+
374  done
375
376
377subsection \<open>Equality rules\<close>
378
379lemma sym: \<open>a = b \<Longrightarrow> b = a\<close>
380  apply (erule subst)
381  apply (rule refl)
382  done
383
384lemma trans: \<open>\<lbrakk>a = b; b = c\<rbrakk> \<Longrightarrow> a = c\<close>
385  apply (erule subst, assumption)
386  done
387
388lemma not_sym: \<open>b \<noteq> a \<Longrightarrow> a \<noteq> b\<close>
389  apply (erule contrapos)
390  apply (erule sym)
391  done
392
393text \<open>
394  Two theorems for rewriting only one instance of a definition:
395  the first for definitions of formulae and the second for terms.
396\<close>
397
398lemma def_imp_iff: \<open>(A \<equiv> B) \<Longrightarrow> A \<longleftrightarrow> B\<close>
399  apply unfold
400  apply (rule iff_refl)
401  done
402
403lemma meta_eq_to_obj_eq: \<open>(A \<equiv> B) \<Longrightarrow> A = B\<close>
404  apply unfold
405  apply (rule refl)
406  done
407
408lemma meta_eq_to_iff: \<open>x \<equiv> y \<Longrightarrow> x \<longleftrightarrow> y\<close>
409  by unfold (rule iff_refl)
410
411text \<open>Substitution.\<close>
412lemma ssubst: \<open>\<lbrakk>b = a; P(a)\<rbrakk> \<Longrightarrow> P(b)\<close>
413  apply (drule sym)
414  apply (erule (1) subst)
415  done
416
417text \<open>A special case of \<open>ex1E\<close> that would otherwise need quantifier
418  expansion.\<close>
419lemma ex1_equalsE: \<open>\<lbrakk>\<exists>!x. P(x); P(a); P(b)\<rbrakk> \<Longrightarrow> a = b\<close>
420  apply (erule ex1E)
421  apply (rule trans)
422   apply (rule_tac [2] sym)
423   apply (assumption | erule spec [THEN mp])+
424  done
425
426
427subsubsection \<open>Polymorphic congruence rules\<close>
428
429lemma subst_context: \<open>a = b \<Longrightarrow> t(a) = t(b)\<close>
430  apply (erule ssubst)
431  apply (rule refl)
432  done
433
434lemma subst_context2: \<open>\<lbrakk>a = b; c = d\<rbrakk> \<Longrightarrow> t(a,c) = t(b,d)\<close>
435  apply (erule ssubst)+
436  apply (rule refl)
437  done
438
439lemma subst_context3: \<open>\<lbrakk>a = b; c = d; e = f\<rbrakk> \<Longrightarrow> t(a,c,e) = t(b,d,f)\<close>
440  apply (erule ssubst)+
441  apply (rule refl)
442  done
443
444text \<open>
445  Useful with \<^ML>\<open>eresolve_tac\<close> for proving equalities from known
446  equalities.
447
448        a = b
449        |   |
450        c = d
451\<close>
452lemma box_equals: \<open>\<lbrakk>a = b; a = c; b = d\<rbrakk> \<Longrightarrow> c = d\<close>
453  apply (rule trans)
454   apply (rule trans)
455    apply (rule sym)
456    apply assumption+
457  done
458
459text \<open>Dual of \<open>box_equals\<close>: for proving equalities backwards.\<close>
460lemma simp_equals: \<open>\<lbrakk>a = c; b = d; c = d\<rbrakk> \<Longrightarrow> a = b\<close>
461  apply (rule trans)
462   apply (rule trans)
463    apply assumption+
464  apply (erule sym)
465  done
466
467
468subsubsection \<open>Congruence rules for predicate letters\<close>
469
470lemma pred1_cong: \<open>a = a' \<Longrightarrow> P(a) \<longleftrightarrow> P(a')\<close>
471  apply (rule iffI)
472   apply (erule (1) subst)
473  apply (erule (1) ssubst)
474  done
475
476lemma pred2_cong: \<open>\<lbrakk>a = a'; b = b'\<rbrakk> \<Longrightarrow> P(a,b) \<longleftrightarrow> P(a',b')\<close>
477  apply (rule iffI)
478   apply (erule subst)+
479   apply assumption
480  apply (erule ssubst)+
481  apply assumption
482  done
483
484lemma pred3_cong: \<open>\<lbrakk>a = a'; b = b'; c = c'\<rbrakk> \<Longrightarrow> P(a,b,c) \<longleftrightarrow> P(a',b',c')\<close>
485  apply (rule iffI)
486   apply (erule subst)+
487   apply assumption
488  apply (erule ssubst)+
489  apply assumption
490  done
491
492text \<open>Special case for the equality predicate!\<close>
493lemma eq_cong: \<open>\<lbrakk>a = a'; b = b'\<rbrakk> \<Longrightarrow> a = b \<longleftrightarrow> a' = b'\<close>
494  apply (erule (1) pred2_cong)
495  done
496
497
498subsection \<open>Simplifications of assumed implications\<close>
499
500text \<open>
501  Roy Dyckhoff has proved that \<open>conj_impE\<close>, \<open>disj_impE\<close>, and
502  \<open>imp_impE\<close> used with \<^ML>\<open>mp_tac\<close> (restricted to atomic formulae) is
503  COMPLETE for intuitionistic propositional logic.
504
505  See R. Dyckhoff, Contraction-free sequent calculi for intuitionistic logic
506  (preprint, University of St Andrews, 1991).
507\<close>
508
509lemma conj_impE:
510  assumes major: \<open>(P \<and> Q) \<longrightarrow> S\<close>
511    and r: \<open>P \<longrightarrow> (Q \<longrightarrow> S) \<Longrightarrow> R\<close>
512  shows \<open>R\<close>
513  by (assumption | rule conjI impI major [THEN mp] r)+
514
515lemma disj_impE:
516  assumes major: \<open>(P \<or> Q) \<longrightarrow> S\<close>
517    and r: \<open>\<lbrakk>P \<longrightarrow> S; Q \<longrightarrow> S\<rbrakk> \<Longrightarrow> R\<close>
518  shows \<open>R\<close>
519  by (assumption | rule disjI1 disjI2 impI major [THEN mp] r)+
520
521text \<open>Simplifies the implication.  Classical version is stronger.
522  Still UNSAFE since Q must be provable -- backtracking needed.\<close>
523lemma imp_impE:
524  assumes major: \<open>(P \<longrightarrow> Q) \<longrightarrow> S\<close>
525    and r1: \<open>\<lbrakk>P; Q \<longrightarrow> S\<rbrakk> \<Longrightarrow> Q\<close>
526    and r2: \<open>S \<Longrightarrow> R\<close>
527  shows \<open>R\<close>
528  by (assumption | rule impI major [THEN mp] r1 r2)+
529
530text \<open>Simplifies the implication.  Classical version is stronger.
531  Still UNSAFE since ~P must be provable -- backtracking needed.\<close>
532lemma not_impE: \<open>\<not> P \<longrightarrow> S \<Longrightarrow> (P \<Longrightarrow> False) \<Longrightarrow> (S \<Longrightarrow> R) \<Longrightarrow> R\<close>
533  apply (drule mp)
534   apply (rule notI)
535   apply assumption
536  apply assumption
537  done
538
539text \<open>Simplifies the implication. UNSAFE.\<close>
540lemma iff_impE:
541  assumes major: \<open>(P \<longleftrightarrow> Q) \<longrightarrow> S\<close>
542    and r1: \<open>\<lbrakk>P; Q \<longrightarrow> S\<rbrakk> \<Longrightarrow> Q\<close>
543    and r2: \<open>\<lbrakk>Q; P \<longrightarrow> S\<rbrakk> \<Longrightarrow> P\<close>
544    and r3: \<open>S \<Longrightarrow> R\<close>
545  shows \<open>R\<close>
546  apply (assumption | rule iffI impI major [THEN mp] r1 r2 r3)+
547  done
548
549text \<open>What if \<open>(\<forall>x. \<not> \<not> P(x)) \<longrightarrow> \<not> \<not> (\<forall>x. P(x))\<close> is an assumption?
550  UNSAFE.\<close>
551lemma all_impE:
552  assumes major: \<open>(\<forall>x. P(x)) \<longrightarrow> S\<close>
553    and r1: \<open>\<And>x. P(x)\<close>
554    and r2: \<open>S \<Longrightarrow> R\<close>
555  shows \<open>R\<close>
556  apply (rule allI impI major [THEN mp] r1 r2)+
557  done
558
559text \<open>
560  Unsafe: \<open>\<exists>x. P(x)) \<longrightarrow> S\<close> is equivalent
561  to \<open>\<forall>x. P(x) \<longrightarrow> S\<close>.\<close>
562lemma ex_impE:
563  assumes major: \<open>(\<exists>x. P(x)) \<longrightarrow> S\<close>
564    and r: \<open>P(x) \<longrightarrow> S \<Longrightarrow> R\<close>
565  shows \<open>R\<close>
566  apply (assumption | rule exI impI major [THEN mp] r)+
567  done
568
569text \<open>Courtesy of Krzysztof Grabczewski.\<close>
570lemma disj_imp_disj: \<open>P \<or> Q \<Longrightarrow> (P \<Longrightarrow> R) \<Longrightarrow> (Q \<Longrightarrow> S) \<Longrightarrow> R \<or> S\<close>
571  apply (erule disjE)
572  apply (rule disjI1) apply assumption
573  apply (rule disjI2) apply assumption
574  done
575
576ML \<open>
577structure Project_Rule = Project_Rule
578(
579  val conjunct1 = @{thm conjunct1}
580  val conjunct2 = @{thm conjunct2}
581  val mp = @{thm mp}
582)
583\<close>
584
585ML_file \<open>fologic.ML\<close>
586
587lemma thin_refl: \<open>\<lbrakk>x = x; PROP W\<rbrakk> \<Longrightarrow> PROP W\<close> .
588
589ML \<open>
590structure Hypsubst = Hypsubst
591(
592  val dest_eq = FOLogic.dest_eq
593  val dest_Trueprop = FOLogic.dest_Trueprop
594  val dest_imp = FOLogic.dest_imp
595  val eq_reflection = @{thm eq_reflection}
596  val rev_eq_reflection = @{thm meta_eq_to_obj_eq}
597  val imp_intr = @{thm impI}
598  val rev_mp = @{thm rev_mp}
599  val subst = @{thm subst}
600  val sym = @{thm sym}
601  val thin_refl = @{thm thin_refl}
602);
603open Hypsubst;
604\<close>
605
606ML_file \<open>intprover.ML\<close>
607
608
609subsection \<open>Intuitionistic Reasoning\<close>
610
611setup \<open>Intuitionistic.method_setup \<^binding>\<open>iprover\<close>\<close>
612
613lemma impE':
614  assumes 1: \<open>P \<longrightarrow> Q\<close>
615    and 2: \<open>Q \<Longrightarrow> R\<close>
616    and 3: \<open>P \<longrightarrow> Q \<Longrightarrow> P\<close>
617  shows \<open>R\<close>
618proof -
619  from 3 and 1 have \<open>P\<close> .
620  with 1 have \<open>Q\<close> by (rule impE)
621  with 2 show \<open>R\<close> .
622qed
623
624lemma allE':
625  assumes 1: \<open>\<forall>x. P(x)\<close>
626    and 2: \<open>P(x) \<Longrightarrow> \<forall>x. P(x) \<Longrightarrow> Q\<close>
627  shows \<open>Q\<close>
628proof -
629  from 1 have \<open>P(x)\<close> by (rule spec)
630  from this and 1 show \<open>Q\<close> by (rule 2)
631qed
632
633lemma notE':
634  assumes 1: \<open>\<not> P\<close>
635    and 2: \<open>\<not> P \<Longrightarrow> P\<close>
636  shows \<open>R\<close>
637proof -
638  from 2 and 1 have \<open>P\<close> .
639  with 1 show \<open>R\<close> by (rule notE)
640qed
641
642lemmas [Pure.elim!] = disjE iffE FalseE conjE exE
643  and [Pure.intro!] = iffI conjI impI TrueI notI allI refl
644  and [Pure.elim 2] = allE notE' impE'
645  and [Pure.intro] = exI disjI2 disjI1
646
647setup \<open>
648  Context_Rules.addSWrapper
649    (fn ctxt => fn tac => hyp_subst_tac ctxt ORELSE' tac)
650\<close>
651
652
653lemma iff_not_sym: \<open>\<not> (Q \<longleftrightarrow> P) \<Longrightarrow> \<not> (P \<longleftrightarrow> Q)\<close>
654  by iprover
655
656lemmas [sym] = sym iff_sym not_sym iff_not_sym
657  and [Pure.elim?] = iffD1 iffD2 impE
658
659
660lemma eq_commute: \<open>a = b \<longleftrightarrow> b = a\<close>
661  apply (rule iffI)
662  apply (erule sym)+
663  done
664
665
666subsection \<open>Atomizing meta-level rules\<close>
667
668lemma atomize_all [atomize]: \<open>(\<And>x. P(x)) \<equiv> Trueprop (\<forall>x. P(x))\<close>
669proof
670  assume \<open>\<And>x. P(x)\<close>
671  then show \<open>\<forall>x. P(x)\<close> ..
672next
673  assume \<open>\<forall>x. P(x)\<close>
674  then show \<open>\<And>x. P(x)\<close> ..
675qed
676
677lemma atomize_imp [atomize]: \<open>(A \<Longrightarrow> B) \<equiv> Trueprop (A \<longrightarrow> B)\<close>
678proof
679  assume \<open>A \<Longrightarrow> B\<close>
680  then show \<open>A \<longrightarrow> B\<close> ..
681next
682  assume \<open>A \<longrightarrow> B\<close> and \<open>A\<close>
683  then show \<open>B\<close> by (rule mp)
684qed
685
686lemma atomize_eq [atomize]: \<open>(x \<equiv> y) \<equiv> Trueprop (x = y)\<close>
687proof
688  assume \<open>x \<equiv> y\<close>
689  show \<open>x = y\<close> unfolding \<open>x \<equiv> y\<close> by (rule refl)
690next
691  assume \<open>x = y\<close>
692  then show \<open>x \<equiv> y\<close> by (rule eq_reflection)
693qed
694
695lemma atomize_iff [atomize]: \<open>(A \<equiv> B) \<equiv> Trueprop (A \<longleftrightarrow> B)\<close>
696proof
697  assume \<open>A \<equiv> B\<close>
698  show \<open>A \<longleftrightarrow> B\<close> unfolding \<open>A \<equiv> B\<close> by (rule iff_refl)
699next
700  assume \<open>A \<longleftrightarrow> B\<close>
701  then show \<open>A \<equiv> B\<close> by (rule iff_reflection)
702qed
703
704lemma atomize_conj [atomize]: \<open>(A &&& B) \<equiv> Trueprop (A \<and> B)\<close>
705proof
706  assume conj: \<open>A &&& B\<close>
707  show \<open>A \<and> B\<close>
708  proof (rule conjI)
709    from conj show \<open>A\<close> by (rule conjunctionD1)
710    from conj show \<open>B\<close> by (rule conjunctionD2)
711  qed
712next
713  assume conj: \<open>A \<and> B\<close>
714  show \<open>A &&& B\<close>
715  proof -
716    from conj show \<open>A\<close> ..
717    from conj show \<open>B\<close> ..
718  qed
719qed
720
721lemmas [symmetric, rulify] = atomize_all atomize_imp
722  and [symmetric, defn] = atomize_all atomize_imp atomize_eq atomize_iff
723
724
725subsection \<open>Atomizing elimination rules\<close>
726
727lemma atomize_exL[atomize_elim]: \<open>(\<And>x. P(x) \<Longrightarrow> Q) \<equiv> ((\<exists>x. P(x)) \<Longrightarrow> Q)\<close>
728  by rule iprover+
729
730lemma atomize_conjL[atomize_elim]: \<open>(A \<Longrightarrow> B \<Longrightarrow> C) \<equiv> (A \<and> B \<Longrightarrow> C)\<close>
731  by rule iprover+
732
733lemma atomize_disjL[atomize_elim]: \<open>((A \<Longrightarrow> C) \<Longrightarrow> (B \<Longrightarrow> C) \<Longrightarrow> C) \<equiv> ((A \<or> B \<Longrightarrow> C) \<Longrightarrow> C)\<close>
734  by rule iprover+
735
736lemma atomize_elimL[atomize_elim]: \<open>(\<And>B. (A \<Longrightarrow> B) \<Longrightarrow> B) \<equiv> Trueprop(A)\<close> ..
737
738
739subsection \<open>Calculational rules\<close>
740
741lemma forw_subst: \<open>a = b \<Longrightarrow> P(b) \<Longrightarrow> P(a)\<close>
742  by (rule ssubst)
743
744lemma back_subst: \<open>P(a) \<Longrightarrow> a = b \<Longrightarrow> P(b)\<close>
745  by (rule subst)
746
747text \<open>
748  Note that this list of rules is in reverse order of priorities.
749\<close>
750
751lemmas basic_trans_rules [trans] =
752  forw_subst
753  back_subst
754  rev_mp
755  mp
756  trans
757
758
759subsection \<open>``Let'' declarations\<close>
760
761nonterminal letbinds and letbind
762
763definition Let :: \<open>['a::{}, 'a => 'b] \<Rightarrow> ('b::{})\<close>
764  where \<open>Let(s, f) \<equiv> f(s)\<close>
765
766syntax
767  "_bind"       :: \<open>[pttrn, 'a] => letbind\<close>           (\<open>(2_ =/ _)\<close> 10)
768  ""            :: \<open>letbind => letbinds\<close>              (\<open>_\<close>)
769  "_binds"      :: \<open>[letbind, letbinds] => letbinds\<close>  (\<open>_;/ _\<close>)
770  "_Let"        :: \<open>[letbinds, 'a] => 'a\<close>             (\<open>(let (_)/ in (_))\<close> 10)
771
772translations
773  "_Let(_binds(b, bs), e)"  == "_Let(b, _Let(bs, e))"
774  "let x = a in e"          == "CONST Let(a, \<lambda>x. e)"
775
776lemma LetI:
777  assumes \<open>\<And>x. x = t \<Longrightarrow> P(u(x))\<close>
778  shows \<open>P(let x = t in u(x))\<close>
779  apply (unfold Let_def)
780  apply (rule refl [THEN assms])
781  done
782
783
784subsection \<open>Intuitionistic simplification rules\<close>
785
786lemma conj_simps:
787  \<open>P \<and> True \<longleftrightarrow> P\<close>
788  \<open>True \<and> P \<longleftrightarrow> P\<close>
789  \<open>P \<and> False \<longleftrightarrow> False\<close>
790  \<open>False \<and> P \<longleftrightarrow> False\<close>
791  \<open>P \<and> P \<longleftrightarrow> P\<close>
792  \<open>P \<and> P \<and> Q \<longleftrightarrow> P \<and> Q\<close>
793  \<open>P \<and> \<not> P \<longleftrightarrow> False\<close>
794  \<open>\<not> P \<and> P \<longleftrightarrow> False\<close>
795  \<open>(P \<and> Q) \<and> R \<longleftrightarrow> P \<and> (Q \<and> R)\<close>
796  by iprover+
797
798lemma disj_simps:
799  \<open>P \<or> True \<longleftrightarrow> True\<close>
800  \<open>True \<or> P \<longleftrightarrow> True\<close>
801  \<open>P \<or> False \<longleftrightarrow> P\<close>
802  \<open>False \<or> P \<longleftrightarrow> P\<close>
803  \<open>P \<or> P \<longleftrightarrow> P\<close>
804  \<open>P \<or> P \<or> Q \<longleftrightarrow> P \<or> Q\<close>
805  \<open>(P \<or> Q) \<or> R \<longleftrightarrow> P \<or> (Q \<or> R)\<close>
806  by iprover+
807
808lemma not_simps:
809  \<open>\<not> (P \<or> Q) \<longleftrightarrow> \<not> P \<and> \<not> Q\<close>
810  \<open>\<not> False \<longleftrightarrow> True\<close>
811  \<open>\<not> True \<longleftrightarrow> False\<close>
812  by iprover+
813
814lemma imp_simps:
815  \<open>(P \<longrightarrow> False) \<longleftrightarrow> \<not> P\<close>
816  \<open>(P \<longrightarrow> True) \<longleftrightarrow> True\<close>
817  \<open>(False \<longrightarrow> P) \<longleftrightarrow> True\<close>
818  \<open>(True \<longrightarrow> P) \<longleftrightarrow> P\<close>
819  \<open>(P \<longrightarrow> P) \<longleftrightarrow> True\<close>
820  \<open>(P \<longrightarrow> \<not> P) \<longleftrightarrow> \<not> P\<close>
821  by iprover+
822
823lemma iff_simps:
824  \<open>(True \<longleftrightarrow> P) \<longleftrightarrow> P\<close>
825  \<open>(P \<longleftrightarrow> True) \<longleftrightarrow> P\<close>
826  \<open>(P \<longleftrightarrow> P) \<longleftrightarrow> True\<close>
827  \<open>(False \<longleftrightarrow> P) \<longleftrightarrow> \<not> P\<close>
828  \<open>(P \<longleftrightarrow> False) \<longleftrightarrow> \<not> P\<close>
829  by iprover+
830
831text \<open>The \<open>x = t\<close> versions are needed for the simplification
832  procedures.\<close>
833lemma quant_simps:
834  \<open>\<And>P. (\<forall>x. P) \<longleftrightarrow> P\<close>
835  \<open>(\<forall>x. x = t \<longrightarrow> P(x)) \<longleftrightarrow> P(t)\<close>
836  \<open>(\<forall>x. t = x \<longrightarrow> P(x)) \<longleftrightarrow> P(t)\<close>
837  \<open>\<And>P. (\<exists>x. P) \<longleftrightarrow> P\<close>
838  \<open>\<exists>x. x = t\<close>
839  \<open>\<exists>x. t = x\<close>
840  \<open>(\<exists>x. x = t \<and> P(x)) \<longleftrightarrow> P(t)\<close>
841  \<open>(\<exists>x. t = x \<and> P(x)) \<longleftrightarrow> P(t)\<close>
842  by iprover+
843
844text \<open>These are NOT supplied by default!\<close>
845lemma distrib_simps:
846  \<open>P \<and> (Q \<or> R) \<longleftrightarrow> P \<and> Q \<or> P \<and> R\<close>
847  \<open>(Q \<or> R) \<and> P \<longleftrightarrow> Q \<and> P \<or> R \<and> P\<close>
848  \<open>(P \<or> Q \<longrightarrow> R) \<longleftrightarrow> (P \<longrightarrow> R) \<and> (Q \<longrightarrow> R)\<close>
849  by iprover+
850
851
852subsubsection \<open>Conversion into rewrite rules\<close>
853
854lemma P_iff_F: \<open>\<not> P \<Longrightarrow> (P \<longleftrightarrow> False)\<close>
855  by iprover
856lemma iff_reflection_F: \<open>\<not> P \<Longrightarrow> (P \<equiv> False)\<close>
857  by (rule P_iff_F [THEN iff_reflection])
858
859lemma P_iff_T: \<open>P \<Longrightarrow> (P \<longleftrightarrow> True)\<close>
860  by iprover
861lemma iff_reflection_T: \<open>P \<Longrightarrow> (P \<equiv> True)\<close>
862  by (rule P_iff_T [THEN iff_reflection])
863
864
865subsubsection \<open>More rewrite rules\<close>
866
867lemma conj_commute: \<open>P \<and> Q \<longleftrightarrow> Q \<and> P\<close> by iprover
868lemma conj_left_commute: \<open>P \<and> (Q \<and> R) \<longleftrightarrow> Q \<and> (P \<and> R)\<close> by iprover
869lemmas conj_comms = conj_commute conj_left_commute
870
871lemma disj_commute: \<open>P \<or> Q \<longleftrightarrow> Q \<or> P\<close> by iprover
872lemma disj_left_commute: \<open>P \<or> (Q \<or> R) \<longleftrightarrow> Q \<or> (P \<or> R)\<close> by iprover
873lemmas disj_comms = disj_commute disj_left_commute
874
875lemma conj_disj_distribL: \<open>P \<and> (Q \<or> R) \<longleftrightarrow> (P \<and> Q \<or> P \<and> R)\<close> by iprover
876lemma conj_disj_distribR: \<open>(P \<or> Q) \<and> R \<longleftrightarrow> (P \<and> R \<or> Q \<and> R)\<close> by iprover
877
878lemma disj_conj_distribL: \<open>P \<or> (Q \<and> R) \<longleftrightarrow> (P \<or> Q) \<and> (P \<or> R)\<close> by iprover
879lemma disj_conj_distribR: \<open>(P \<and> Q) \<or> R \<longleftrightarrow> (P \<or> R) \<and> (Q \<or> R)\<close> by iprover
880
881lemma imp_conj_distrib: \<open>(P \<longrightarrow> (Q \<and> R)) \<longleftrightarrow> (P \<longrightarrow> Q) \<and> (P \<longrightarrow> R)\<close> by iprover
882lemma imp_conj: \<open>((P \<and> Q) \<longrightarrow> R) \<longleftrightarrow> (P \<longrightarrow> (Q \<longrightarrow> R))\<close> by iprover
883lemma imp_disj: \<open>(P \<or> Q \<longrightarrow> R) \<longleftrightarrow> (P \<longrightarrow> R) \<and> (Q \<longrightarrow> R)\<close> by iprover
884
885lemma de_Morgan_disj: \<open>(\<not> (P \<or> Q)) \<longleftrightarrow> (\<not> P \<and> \<not> Q)\<close> by iprover
886
887lemma not_ex: \<open>(\<not> (\<exists>x. P(x))) \<longleftrightarrow> (\<forall>x. \<not> P(x))\<close> by iprover
888lemma imp_ex: \<open>((\<exists>x. P(x)) \<longrightarrow> Q) \<longleftrightarrow> (\<forall>x. P(x) \<longrightarrow> Q)\<close> by iprover
889
890lemma ex_disj_distrib: \<open>(\<exists>x. P(x) \<or> Q(x)) \<longleftrightarrow> ((\<exists>x. P(x)) \<or> (\<exists>x. Q(x)))\<close>
891  by iprover
892
893lemma all_conj_distrib: \<open>(\<forall>x. P(x) \<and> Q(x)) \<longleftrightarrow> ((\<forall>x. P(x)) \<and> (\<forall>x. Q(x)))\<close>
894  by iprover
895
896end
897