1 (* Title: FOLP/IFOLP.thy
2 Author: Martin D Coen, Cambridge University Computer Laboratory
3 Copyright 1992 University of Cambridge
6 header {* Intuitionistic First-Order Logic with Proofs *}
12 ML_file "~~/src/Tools/misc_legacy.ML"
14 setup Pure_Thy.old_appl_syntax_setup
24 Proof :: "[o,p]=>prop"
25 EqProof :: "[p,p,o]=>prop" ("(3_ /= _ :/ _)" [10,10,10] 5)
27 (*** Logical Connectives -- Type Formers ***)
28 eq :: "['a,'a] => o" (infixl "=" 50)
31 Not :: "o => o" ("~ _" [40] 40)
32 conj :: "[o,o] => o" (infixr "&" 35)
33 disj :: "[o,o] => o" (infixr "|" 30)
34 imp :: "[o,o] => o" (infixr "-->" 25)
35 iff :: "[o,o] => o" (infixr "<->" 25)
37 All :: "('a => o) => o" (binder "ALL " 10)
38 Ex :: "('a => o) => o" (binder "EX " 10)
39 Ex1 :: "('a => o) => o" (binder "EX! " 10)
44 (*** Proof Term Formers: precedence must exceed 50 ***)
49 pair :: "[p,p]=>p" ("(1<_,/_>)")
50 split :: "[p, [p,p]=>p] =>p"
53 when :: "[p, p=>p, p=>p]=>p"
54 lambda :: "(p => p) => p" (binder "lam " 55)
55 App :: "[p,p]=>p" (infixl "`" 60)
56 alll :: "['a=>p]=>p" (binder "all " 55)
57 app :: "[p,'a]=>p" (infixl "^" 55)
58 exists :: "['a,p]=>p" ("(1[_,/_])")
59 xsplit :: "[p,['a,p]=>p]=>p"
61 idpeel :: "[p,'a=>p]=>p"
65 syntax "_Proof" :: "[p,o]=>prop" ("(_ /: _)" [51, 10] 5)
68 let fun proof_tr [p, P] = Const (@{const_syntax Proof}, dummyT) $ P $ p
69 in [(@{syntax_const "_Proof"}, proof_tr)] end
72 (*show_proofs = true displays the proof terms -- they are ENORMOUS*)
73 ML {* val show_proofs = Attrib.setup_config_bool @{binding show_proofs} (K false) *}
75 print_translation (advanced) {*
77 fun proof_tr' ctxt [P, p] =
78 if Config.get ctxt show_proofs then Const (@{syntax_const "_Proof"}, dummyT) $ p $ P
80 in [(@{const_syntax Proof}, proof_tr')] end
85 (**** Propositional logic ****)
88 (* Like Intensional Equality in MLTT - but proofs distinct from terms *)
91 ieqE: "[| p : a=b; !!x. f(x) : P(x,x) |] ==> idpeel(p,f) : P(a,b)"
93 (* Truth and Falsity *)
96 FalseE: "a:False ==> contr(a):P"
100 conjI: "[| a:P; b:Q |] ==> <a,b> : P&Q"
101 conjunct1: "p:P&Q ==> fst(p):P"
102 conjunct2: "p:P&Q ==> snd(p):Q"
106 disjI1: "a:P ==> inl(a):P|Q"
107 disjI2: "b:Q ==> inr(b):P|Q"
108 disjE: "[| a:P|Q; !!x. x:P ==> f(x):R; !!x. x:Q ==> g(x):R
109 |] ==> when(a,f,g):R"
113 impI: "(!!x. x:P ==> f(x):Q) ==> lam x. f(x):P-->Q"
114 mp: "[| f:P-->Q; a:P |] ==> f`a:Q"
118 allI: "(!!x. f(x) : P(x)) ==> all x. f(x) : ALL x. P(x)"
119 spec: "(f:ALL x. P(x)) ==> f^x : P(x)"
121 exI: "p : P(x) ==> [x,p] : EX x. P(x)"
122 exE: "[| p: EX x. P(x); !!x u. u:P(x) ==> f(x,u) : R |] ==> xsplit(p,f):R"
124 (**** Equality between proofs ****)
126 prefl: "a : P ==> a = a : P"
127 psym: "a = b : P ==> b = a : P"
128 ptrans: "[| a = b : P; b = c : P |] ==> a = c : P"
130 idpeelB: "[| !!x. f(x) : P(x,x) |] ==> idpeel(ideq(a),f) = f(a) : P(a,a)"
132 fstB: "a:P ==> fst(<a,b>) = a : P"
133 sndB: "b:Q ==> snd(<a,b>) = b : Q"
134 pairEC: "p:P&Q ==> p = <fst(p),snd(p)> : P&Q"
136 whenBinl: "[| a:P; !!x. x:P ==> f(x) : Q |] ==> when(inl(a),f,g) = f(a) : Q"
137 whenBinr: "[| b:P; !!x. x:P ==> g(x) : Q |] ==> when(inr(b),f,g) = g(b) : Q"
138 plusEC: "a:P|Q ==> when(a,%x. inl(x),%y. inr(y)) = a : P|Q"
140 applyB: "[| a:P; !!x. x:P ==> b(x) : Q |] ==> (lam x. b(x)) ` a = b(a) : Q"
141 funEC: "f:P ==> f = lam x. f`x : P"
143 specB: "[| !!x. f(x) : P(x) |] ==> (all x. f(x)) ^ a = f(a) : P(a)"
146 (**** Definitions ****)
148 not_def: "~P == P-->False"
149 iff_def: "P<->Q == (P-->Q) & (Q-->P)"
152 ex1_def: "EX! x. P(x) == EX x. P(x) & (ALL y. P(y) --> y=x)"
154 (*Rewriting -- special constants to flag normalized terms and formulae*)
155 norm_eq: "nrm : norm(x) = x"
156 NORM_iff: "NRM : NORM(P) <-> P"
158 (*** Sequent-style elimination rules for & --> and ALL ***)
160 schematic_lemma conjE:
162 and "!!x y.[| x:P; y:Q |] ==> f(x,y):R"
164 apply (rule assms(2))
165 apply (rule conjunct1 [OF assms(1)])
166 apply (rule conjunct2 [OF assms(1)])
169 schematic_lemma impE:
172 and "!!x. x:Q ==> r(x):R"
174 apply (rule assms mp)+
177 schematic_lemma allE:
178 assumes "p:ALL x. P(x)"
179 and "!!y. y:P(x) ==> q(y):R"
181 apply (rule assms spec)+
184 (*Duplicates the quantifier; for use with eresolve_tac*)
185 schematic_lemma all_dupE:
186 assumes "p:ALL x. P(x)"
187 and "!!y z.[| y:P(x); z:ALL x. P(x) |] ==> q(y,z):R"
189 apply (rule assms spec)+
193 (*** Negation rules, which translate between ~P and P-->False ***)
195 schematic_lemma notI:
196 assumes "!!x. x:P ==> q(x):False"
199 apply (assumption | rule assms impI)+
202 schematic_lemma notE: "p:~P \<Longrightarrow> q:P \<Longrightarrow> ?p:R"
208 (*This is useful with the special implication rules for each kind of P. *)
209 schematic_lemma not_to_imp:
211 and "!!x. x:(P-->False) ==> q(x):Q"
213 apply (assumption | rule assms impI notE)+
216 (* For substitution int an assumption P, reduce Q to P-->Q, substitute into
217 this implication, then apply impI to move P back into the assumptions.*)
218 schematic_lemma rev_mp: "[| p:P; q:P --> Q |] ==> ?p:Q"
219 apply (assumption | rule mp)+
223 (*Contrapositive of an inference rule*)
224 schematic_lemma contrapos:
225 assumes major: "p:~Q"
226 and minor: "!!y. y:P==>q(y):Q"
228 apply (rule major [THEN notE, THEN notI])
232 (** Unique assumption tactic.
233 Ignores proof objects.
234 Fails unless one assumption is equal and exactly one is unifiable
239 fun discard_proof (Const (@{const_name Proof}, _) $ P $ _) = P;
241 val uniq_assume_tac =
244 let val hyps = map discard_proof (Logic.strip_assums_hyp prem)
245 and concl = discard_proof (Logic.strip_assums_concl prem)
247 if exists (fn hyp => hyp aconv concl) hyps
248 then case distinct (op =) (filter (fn hyp => Term.could_unify (hyp, concl)) hyps) of
257 (*** Modus Ponens Tactics ***)
259 (*Finds P-->Q and P in the assumptions, replaces implication by Q *)
261 fun mp_tac i = eresolve_tac [@{thm notE}, make_elim @{thm mp}] i THEN assume_tac i
264 (*Like mp_tac but instantiates no variables*)
266 fun int_uniq_mp_tac i = eresolve_tac [@{thm notE}, @{thm impE}] i THEN uniq_assume_tac i
270 (*** If-and-only-if ***)
272 schematic_lemma iffI:
273 assumes "!!x. x:P ==> q(x):Q"
274 and "!!x. x:Q ==> r(x):P"
277 apply (assumption | rule assms conjI impI)+
281 (*Observe use of rewrite_rule to unfold "<->" in meta-assumptions (prems) *)
283 schematic_lemma iffE:
285 and "!!x y.[| x:P-->Q; y:Q-->P |] ==> q(x,y):R"
288 apply (rule assms(1) [unfolded iff_def])
289 apply (rule assms(2))
293 (* Destruct rules for <-> similar to Modus Ponens *)
295 schematic_lemma iffD1: "[| p:P <-> Q; q:P |] ==> ?p:Q"
297 apply (rule conjunct1 [THEN mp], assumption+)
300 schematic_lemma iffD2: "[| p:P <-> Q; q:Q |] ==> ?p:P"
302 apply (rule conjunct2 [THEN mp], assumption+)
305 schematic_lemma iff_refl: "?p:P <-> P"
310 schematic_lemma iff_sym: "p:Q <-> P ==> ?p:P <-> Q"
313 apply (erule (1) mp)+
316 schematic_lemma iff_trans: "[| p:P <-> Q; q:Q<-> R |] ==> ?p:P <-> R"
318 apply (assumption | erule iffE | erule (1) impE)+
321 (*** Unique existence. NOTE THAT the following 2 quantifications
322 EX!x such that [EX!y such that P(x,y)] (sequential)
323 EX!x,y such that P(x,y) (simultaneous)
324 do NOT mean the same thing. The parser treats EX!x y.P(x,y) as sequential.
327 schematic_lemma ex1I:
329 and "!!x u. u:P(x) ==> f(u) : x=a"
330 shows "?p:EX! x. P(x)"
332 apply (assumption | rule assms exI conjI allI impI)+
335 schematic_lemma ex1E:
336 assumes "p:EX! x. P(x)"
337 and "!!x u v. [| u:P(x); v:ALL y. P(y) --> y=x |] ==> f(x,u,v):R"
339 apply (insert assms(1) [unfolded ex1_def])
340 apply (erule exE conjE | assumption | rule assms(1))+
341 apply (erule assms(2), assumption)
345 (*** <-> congruence rules for simplification ***)
347 (*Use iffE on a premise. For conj_cong, imp_cong, all_cong, ex_cong*)
349 fun iff_tac prems i =
350 resolve_tac (prems RL [@{thm iffE}]) i THEN
351 REPEAT1 (eresolve_tac [asm_rl, @{thm mp}] i)
354 schematic_lemma conj_cong:
356 and "!!x. x:P' ==> q(x):Q <-> Q'"
357 shows "?p:(P&Q) <-> (P'&Q')"
358 apply (insert assms(1))
359 apply (assumption | rule iffI conjI |
360 erule iffE conjE mp | tactic {* iff_tac @{thms assms} 1 *})+
363 schematic_lemma disj_cong:
364 "[| p:P <-> P'; q:Q <-> Q' |] ==> ?p:(P|Q) <-> (P'|Q')"
365 apply (erule iffE disjE disjI1 disjI2 | assumption | rule iffI | tactic {* mp_tac 1 *})+
368 schematic_lemma imp_cong:
370 and "!!x. x:P' ==> q(x):Q <-> Q'"
371 shows "?p:(P-->Q) <-> (P'-->Q')"
372 apply (insert assms(1))
373 apply (assumption | rule iffI impI | erule iffE | tactic {* mp_tac 1 *} |
374 tactic {* iff_tac @{thms assms} 1 *})+
377 schematic_lemma iff_cong:
378 "[| p:P <-> P'; q:Q <-> Q' |] ==> ?p:(P<->Q) <-> (P'<->Q')"
379 apply (erule iffE | assumption | rule iffI | tactic {* mp_tac 1 *})+
382 schematic_lemma not_cong:
383 "p:P <-> P' ==> ?p:~P <-> ~P'"
384 apply (assumption | rule iffI notI | tactic {* mp_tac 1 *} | erule iffE notE)+
387 schematic_lemma all_cong:
388 assumes "!!x. f(x):P(x) <-> Q(x)"
389 shows "?p:(ALL x. P(x)) <-> (ALL x. Q(x))"
390 apply (assumption | rule iffI allI | tactic {* mp_tac 1 *} | erule allE |
391 tactic {* iff_tac @{thms assms} 1 *})+
394 schematic_lemma ex_cong:
395 assumes "!!x. f(x):P(x) <-> Q(x)"
396 shows "?p:(EX x. P(x)) <-> (EX x. Q(x))"
397 apply (erule exE | assumption | rule iffI exI | tactic {* mp_tac 1 *} |
398 tactic {* iff_tac @{thms assms} 1 *})+
402 bind_thm ("ex1_cong", prove_goal (the_context ())
403 "(!!x.f(x):P(x) <-> Q(x)) ==> ?p:(EX! x.P(x)) <-> (EX! x.Q(x))"
405 [ (REPEAT (eresolve_tac [ex1E, spec RS mp] 1 ORELSE ares_tac [iffI,ex1I] 1
407 ORELSE iff_tac prems 1)) ]))
410 (*** Equality rules ***)
414 schematic_lemma subst:
415 assumes prem1: "p:a=b"
418 apply (rule prem2 [THEN rev_mp])
419 apply (rule prem1 [THEN ieqE])
424 schematic_lemma sym: "q:a=b ==> ?c:b=a"
429 schematic_lemma trans: "[| p:a=b; q:b=c |] ==> ?d:a=c"
430 apply (erule (1) subst)
433 (** ~ b=a ==> ~ a=b **)
434 schematic_lemma not_sym: "p:~ b=a ==> ?q:~ a=b"
435 apply (erule contrapos)
439 schematic_lemma ssubst: "p:b=a \<Longrightarrow> q:P(a) \<Longrightarrow> ?p:P(b)"
445 (*A special case of ex1E that would otherwise need quantifier expansion*)
446 schematic_lemma ex1_equalsE: "[| p:EX! x. P(x); q:P(a); r:P(b) |] ==> ?d:a=b"
449 apply (rule_tac [2] sym)
450 apply (assumption | erule spec [THEN mp])+
453 (** Polymorphic congruence rules **)
455 schematic_lemma subst_context: "[| p:a=b |] ==> ?d:t(a)=t(b)"
460 schematic_lemma subst_context2: "[| p:a=b; q:c=d |] ==> ?p:t(a,c)=t(b,d)"
461 apply (erule ssubst)+
465 schematic_lemma subst_context3: "[| p:a=b; q:c=d; r:e=f |] ==> ?p:t(a,c,e)=t(b,d,f)"
466 apply (erule ssubst)+
470 (*Useful with eresolve_tac for proving equalties from known equalities.
474 schematic_lemma box_equals: "[| p:a=b; q:a=c; r:b=d |] ==> ?p:c=d"
481 (*Dual of box_equals: for proving equalities backwards*)
482 schematic_lemma simp_equals: "[| p:a=c; q:b=d; r:c=d |] ==> ?p:a=b"
485 apply (assumption | rule sym)+
488 (** Congruence rules for predicate letters **)
490 schematic_lemma pred1_cong: "p:a=a' ==> ?p:P(a) <-> P(a')"
492 apply (tactic {* DEPTH_SOLVE (atac 1 ORELSE eresolve_tac [@{thm subst}, @{thm ssubst}] 1) *})
495 schematic_lemma pred2_cong: "[| p:a=a'; q:b=b' |] ==> ?p:P(a,b) <-> P(a',b')"
497 apply (tactic {* DEPTH_SOLVE (atac 1 ORELSE eresolve_tac [@{thm subst}, @{thm ssubst}] 1) *})
500 schematic_lemma pred3_cong: "[| p:a=a'; q:b=b'; r:c=c' |] ==> ?p:P(a,b,c) <-> P(a',b',c')"
502 apply (tactic {* DEPTH_SOLVE (atac 1 ORELSE eresolve_tac [@{thm subst}, @{thm ssubst}] 1) *})
505 lemmas pred_congs = pred1_cong pred2_cong pred3_cong
507 (*special case for the equality predicate!*)
508 lemmas eq_cong = pred2_cong [where P = "op ="]
511 (*** Simplifications of assumed implications.
512 Roy Dyckhoff has proved that conj_impE, disj_impE, and imp_impE
513 used with mp_tac (restricted to atomic formulae) is COMPLETE for
514 intuitionistic propositional logic. See
515 R. Dyckhoff, Contraction-free sequent calculi for intuitionistic logic
516 (preprint, University of St Andrews, 1991) ***)
518 schematic_lemma conj_impE:
519 assumes major: "p:(P&Q)-->S"
520 and minor: "!!x. x:P-->(Q-->S) ==> q(x):R"
522 apply (assumption | rule conjI impI major [THEN mp] minor)+
525 schematic_lemma disj_impE:
526 assumes major: "p:(P|Q)-->S"
527 and minor: "!!x y.[| x:P-->S; y:Q-->S |] ==> q(x,y):R"
529 apply (tactic {* DEPTH_SOLVE (atac 1 ORELSE
530 resolve_tac [@{thm disjI1}, @{thm disjI2}, @{thm impI},
531 @{thm major} RS @{thm mp}, @{thm minor}] 1) *})
534 (*Simplifies the implication. Classical version is stronger.
535 Still UNSAFE since Q must be provable -- backtracking needed. *)
536 schematic_lemma imp_impE:
537 assumes major: "p:(P-->Q)-->S"
538 and r1: "!!x y.[| x:P; y:Q-->S |] ==> q(x,y):Q"
539 and r2: "!!x. x:S ==> r(x):R"
541 apply (assumption | rule impI major [THEN mp] r1 r2)+
544 (*Simplifies the implication. Classical version is stronger.
545 Still UNSAFE since ~P must be provable -- backtracking needed. *)
546 schematic_lemma not_impE:
547 assumes major: "p:~P --> S"
548 and r1: "!!y. y:P ==> q(y):False"
549 and r2: "!!y. y:S ==> r(y):R"
551 apply (assumption | rule notI impI major [THEN mp] r1 r2)+
554 (*Simplifies the implication. UNSAFE. *)
555 schematic_lemma iff_impE:
556 assumes major: "p:(P<->Q)-->S"
557 and r1: "!!x y.[| x:P; y:Q-->S |] ==> q(x,y):Q"
558 and r2: "!!x y.[| x:Q; y:P-->S |] ==> r(x,y):P"
559 and r3: "!!x. x:S ==> s(x):R"
561 apply (assumption | rule iffI impI major [THEN mp] r1 r2 r3)+
564 (*What if (ALL x.~~P(x)) --> ~~(ALL x.P(x)) is an assumption? UNSAFE*)
565 schematic_lemma all_impE:
566 assumes major: "p:(ALL x. P(x))-->S"
567 and r1: "!!x. q:P(x)"
568 and r2: "!!y. y:S ==> r(y):R"
570 apply (assumption | rule allI impI major [THEN mp] r1 r2)+
573 (*Unsafe: (EX x.P(x))-->S is equivalent to ALL x.P(x)-->S. *)
574 schematic_lemma ex_impE:
575 assumes major: "p:(EX x. P(x))-->S"
576 and r: "!!y. y:P(a)-->S ==> q(y):R"
578 apply (assumption | rule exI impI major [THEN mp] r)+
582 schematic_lemma rev_cut_eq:
584 and "!!x. x:a=b ==> f(x):R"
589 lemma thin_refl: "!!X. [|p:x=x; PROP W|] ==> PROP W" .
591 ML_file "hypsubst.ML"
594 structure Hypsubst = Hypsubst
596 (*Take apart an equality judgement; otherwise raise Match!*)
597 fun dest_eq (Const (@{const_name Proof}, _) $
598 (Const (@{const_name eq}, _) $ t $ u) $ _) = (t, u);
600 val imp_intr = @{thm impI}
602 (*etac rev_cut_eq moves an equality to be the last premise. *)
603 val rev_cut_eq = @{thm rev_cut_eq}
605 val rev_mp = @{thm rev_mp}
606 val subst = @{thm subst}
608 val thin_refl = @{thm thin_refl}
613 ML_file "intprover.ML"
616 (*** Rewrite rules ***)
618 schematic_lemma conj_rews:
619 "?p1 : P & True <-> P"
620 "?p2 : True & P <-> P"
621 "?p3 : P & False <-> False"
622 "?p4 : False & P <-> False"
624 "?p6 : P & ~P <-> False"
625 "?p7 : ~P & P <-> False"
626 "?p8 : (P & Q) & R <-> P & (Q & R)"
627 apply (tactic {* fn st => IntPr.fast_tac 1 st *})+
630 schematic_lemma disj_rews:
631 "?p1 : P | True <-> True"
632 "?p2 : True | P <-> True"
633 "?p3 : P | False <-> P"
634 "?p4 : False | P <-> P"
636 "?p6 : (P | Q) | R <-> P | (Q | R)"
637 apply (tactic {* IntPr.fast_tac 1 *})+
640 schematic_lemma not_rews:
641 "?p1 : ~ False <-> True"
642 "?p2 : ~ True <-> False"
643 apply (tactic {* IntPr.fast_tac 1 *})+
646 schematic_lemma imp_rews:
647 "?p1 : (P --> False) <-> ~P"
648 "?p2 : (P --> True) <-> True"
649 "?p3 : (False --> P) <-> True"
650 "?p4 : (True --> P) <-> P"
651 "?p5 : (P --> P) <-> True"
652 "?p6 : (P --> ~P) <-> ~P"
653 apply (tactic {* IntPr.fast_tac 1 *})+
656 schematic_lemma iff_rews:
657 "?p1 : (True <-> P) <-> P"
658 "?p2 : (P <-> True) <-> P"
659 "?p3 : (P <-> P) <-> True"
660 "?p4 : (False <-> P) <-> ~P"
661 "?p5 : (P <-> False) <-> ~P"
662 apply (tactic {* IntPr.fast_tac 1 *})+
665 schematic_lemma quant_rews:
666 "?p1 : (ALL x. P) <-> P"
667 "?p2 : (EX x. P) <-> P"
668 apply (tactic {* IntPr.fast_tac 1 *})+
671 (*These are NOT supplied by default!*)
672 schematic_lemma distrib_rews1:
673 "?p1 : ~(P|Q) <-> ~P & ~Q"
674 "?p2 : P & (Q | R) <-> P&Q | P&R"
675 "?p3 : (Q | R) & P <-> Q&P | R&P"
676 "?p4 : (P | Q --> R) <-> (P --> R) & (Q --> R)"
677 apply (tactic {* IntPr.fast_tac 1 *})+
680 schematic_lemma distrib_rews2:
681 "?p1 : ~(EX x. NORM(P(x))) <-> (ALL x. ~NORM(P(x)))"
682 "?p2 : ((EX x. NORM(P(x))) --> Q) <-> (ALL x. NORM(P(x)) --> Q)"
683 "?p3 : (EX x. NORM(P(x))) & NORM(Q) <-> (EX x. NORM(P(x)) & NORM(Q))"
684 "?p4 : NORM(Q) & (EX x. NORM(P(x))) <-> (EX x. NORM(Q) & NORM(P(x)))"
685 apply (tactic {* IntPr.fast_tac 1 *})+
688 lemmas distrib_rews = distrib_rews1 distrib_rews2
690 schematic_lemma P_Imp_P_iff_T: "p:P ==> ?p:(P <-> True)"
691 apply (tactic {* IntPr.fast_tac 1 *})
694 schematic_lemma not_P_imp_P_iff_F: "p:~P ==> ?p:(P <-> False)"
695 apply (tactic {* IntPr.fast_tac 1 *})