| author | wenzelm | 
| Wed, 06 Dec 2023 20:16:23 +0100 | |
| changeset 79155 | 53288743c2f0 | 
| parent 76091 | 922e3f9251ac | 
| permissions | -rw-r--r-- | 
| 21425 | 1 | (* Title: HOL/Prolog/prolog.ML | 
| 2 | Author: David von Oheimb (based on a lecture on Lambda Prolog by Nadathur) | |
| 3 | *) | |
| 4 | ||
| 5 | structure Prolog = | |
| 6 | struct | |
| 7 | ||
| 8 | exception not_HOHH; | |
| 9 | ||
| 10 | fun isD t = case t of | |
| 69597 | 11 | Const(\<^const_name>\<open>Trueprop\<close>,_)$t => isD t | 
| 12 | | Const(\<^const_name>\<open>HOL.conj\<close> ,_)$l$r => isD l andalso isD r | |
| 13 | | Const(\<^const_name>\<open>HOL.implies\<close>,_)$l$r => isG l andalso isD r | |
| 14 | | Const(\<^const_name>\<open>Pure.imp\<close>,_)$l$r => isG l andalso isD r | |
| 15 | | Const(\<^const_name>\<open>All\<close>,_)$Abs(s,_,t) => isD t | |
| 16 | | Const(\<^const_name>\<open>Pure.all\<close>,_)$Abs(s,_,t) => isD t | |
| 17 | | Const(\<^const_name>\<open>HOL.disj\<close>,_)$_$_ => false | |
| 18 | | Const(\<^const_name>\<open>Ex\<close> ,_)$_ => false | |
| 19 | | Const(\<^const_name>\<open>Not\<close>,_)$_ => false | |
| 20 | | Const(\<^const_name>\<open>True\<close>,_) => false | |
| 21 | | Const(\<^const_name>\<open>False\<close>,_) => false | |
| 21425 | 22 | | l $ r => isD l | 
| 23 | | Const _ (* rigid atom *) => true | |
| 24 | | Bound _ (* rigid atom *) => true | |
| 25 | | Free _ (* rigid atom *) => true | |
| 26 | | _ (* flexible atom, | |
| 27 | anything else *) => false | |
| 28 | and | |
| 29 | isG t = case t of | |
| 69597 | 30 | Const(\<^const_name>\<open>Trueprop\<close>,_)$t => isG t | 
| 31 | | Const(\<^const_name>\<open>HOL.conj\<close> ,_)$l$r => isG l andalso isG r | |
| 32 | | Const(\<^const_name>\<open>HOL.disj\<close> ,_)$l$r => isG l andalso isG r | |
| 33 | | Const(\<^const_name>\<open>HOL.implies\<close>,_)$l$r => isD l andalso isG r | |
| 34 | | Const(\<^const_name>\<open>Pure.imp\<close>,_)$l$r => isD l andalso isG r | |
| 35 | | Const(\<^const_name>\<open>All\<close>,_)$Abs(_,_,t) => isG t | |
| 36 | | Const(\<^const_name>\<open>Pure.all\<close>,_)$Abs(_,_,t) => isG t | |
| 37 | | Const(\<^const_name>\<open>Ex\<close> ,_)$Abs(_,_,t) => isG t | |
| 38 | | Const(\<^const_name>\<open>True\<close>,_) => true | |
| 39 | | Const(\<^const_name>\<open>Not\<close>,_)$_ => false | |
| 40 | | Const(\<^const_name>\<open>False\<close>,_) => false | |
| 21425 | 41 | | _ (* atom *) => true; | 
| 42 | ||
| 43 | val check_HOHH_tac1 = PRIMITIVE (fn thm => | |
| 59582 | 44 | if isG (Thm.concl_of thm) then thm else raise not_HOHH); | 
| 21425 | 45 | val check_HOHH_tac2 = PRIMITIVE (fn thm => | 
| 59582 | 46 | if forall isG (Thm.prems_of thm) then thm else raise not_HOHH); | 
| 47 | fun check_HOHH thm = (if isD (Thm.concl_of thm) andalso forall isG (Thm.prems_of thm) | |
| 21425 | 48 | then thm else raise not_HOHH); | 
| 49 | ||
| 55143 
04448228381d
explicit eigen-context for attributes "where", "of", and corresponding read_instantiate, instantiate_tac;
 wenzelm parents: 
52233diff
changeset | 50 | fun atomizeD ctxt thm = | 
| 
04448228381d
explicit eigen-context for attributes "where", "of", and corresponding read_instantiate, instantiate_tac;
 wenzelm parents: 
52233diff
changeset | 51 | let | 
| 59582 | 52 | fun at thm = case Thm.concl_of thm of | 
| 69597 | 53 | _$(Const(\<^const_name>\<open>All\<close> ,_)$Abs(s,_,_))=> | 
| 59755 | 54 | let val s' = if s="P" then "PP" else s in | 
| 55 |           at(thm RS (Rule_Insts.read_instantiate ctxt [((("x", 0), Position.none), s')] [s'] spec))
 | |
| 56 | end | |
| 69597 | 57 | | _$(Const(\<^const_name>\<open>HOL.conj\<close>,_)$_$_) => at(thm RS conjunct1)@at(thm RS conjunct2) | 
| 58 | | _$(Const(\<^const_name>\<open>HOL.implies\<close>,_)$_$_) => at(thm RS mp) | |
| 21425 | 59 | | _ => [thm] | 
| 55143 
04448228381d
explicit eigen-context for attributes "where", "of", and corresponding read_instantiate, instantiate_tac;
 wenzelm parents: 
52233diff
changeset | 60 | in map zero_var_indexes (at thm) end; | 
| 21425 | 61 | |
| 62 | val atomize_ss = | |
| 69597 | 63 | (empty_simpset \<^context> |> Simplifier.set_mksimps (mksimps mksimps_pairs)) | 
| 21425 | 64 | addsimps [ | 
| 45654 | 65 |         @{thm all_conj_distrib}, (* "(! x. P x & Q x) = ((! x. P x) & (! x. Q x))" *)
 | 
| 46161 | 66 |         @{thm imp_conjL} RS sym, (* "(D :- G1 :- G2) = (D :- G1 & G2)" *)
 | 
| 67 |         @{thm imp_conjR},        (* "(D1 & D2 :- G) = ((D1 :- G) & (D2 :- G))" *)
 | |
| 52088 | 68 |         @{thm imp_all}]          (* "((!x. D) :- G) = (!x. D :- G)" *)
 | 
| 69 | |> simpset_of; | |
| 21425 | 70 | |
| 45625 
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
39125diff
changeset | 71 | |
| 32283 | 72 | (*val hyp_resolve_tac = Subgoal.FOCUS_PREMS (fn {prems, ...} =>
 | 
| 32260 | 73 | resolve_tac (maps atomizeD prems) 1); | 
| 21425 | 74 | -- is nice, but cannot instantiate unknowns in the assumptions *) | 
| 60754 | 75 | fun hyp_resolve_tac ctxt = SUBGOAL (fn (subgoal, i) => | 
| 46473 | 76 | let | 
| 69597 | 77 | fun ap (Const(\<^const_name>\<open>All\<close>,_)$Abs(_,_,t))=(case ap t of (k,a,t) => (k+1,a ,t)) | 
| 78 | | ap (Const(\<^const_name>\<open>HOL.implies\<close>,_)$_$t) =(case ap t of (k,_,t) => (k,true ,t)) | |
| 21425 | 79 | | ap t = (0,false,t); | 
| 80 | (* | |
| 56245 | 81 |         fun rep_goal (Const (@{const_name Pure.all},_)$Abs (_,_,t)) = rep_goal t
 | 
| 82 |         |   rep_goal (Const (@{const_name Pure.imp},_)$s$t)         =
 | |
| 21425 | 83 | (case rep_goal t of (l,t) => (s::l,t)) | 
| 84 | | rep_goal t = ([] ,t); | |
| 38557 | 85 |         val (prems, Const(@{const_name Trueprop}, _)$concl) = rep_goal
 | 
| 21425 | 86 | (#3(dest_state (st,i))); | 
| 87 | *) | |
| 88 | val prems = Logic.strip_assums_hyp subgoal; | |
| 89 | val concl = HOLogic.dest_Trueprop (Logic.strip_assums_concl subgoal); | |
| 90 | fun drot_tac k i = DETERM (rotate_tac k i); | |
| 91 | fun spec_tac 0 i = all_tac | |
| 60754 | 92 | | spec_tac k i = EVERY' [dresolve_tac ctxt [spec], drot_tac ~1, spec_tac (k-1)] i; | 
| 21425 | 93 | fun dup_spec_tac k i = if k = 0 then all_tac else EVERY' | 
| 60754 | 94 | [DETERM o (eresolve_tac ctxt [all_dupE]), drot_tac ~2, spec_tac (k-1)] i; | 
| 21425 | 95 | fun same_head _ (Const (x,_)) (Const (y,_)) = x = y | 
| 96 | | same_head k (s$_) (t$_) = same_head k s t | |
| 97 | | same_head k (Bound i) (Bound j) = i = j + k | |
| 98 | | same_head _ _ _ = true; | |
| 99 | fun mapn f n [] = [] | |
| 100 | | mapn f n (x::xs) = f n x::mapn f (n+1) xs; | |
| 101 | fun pres_tac (k,arrow,t) n i = drot_tac n i THEN ( | |
| 60754 | 102 | if same_head k t concl | 
| 103 | then dup_spec_tac k i THEN | |
| 104 | (if arrow then eresolve_tac ctxt [mp] i THEN drot_tac (~n) i else assume_tac ctxt i) | |
| 105 | else no_tac); | |
| 21425 | 106 | val ptacs = mapn (fn n => fn t => | 
| 107 | pres_tac (ap (HOLogic.dest_Trueprop t)) n i) 0 prems; | |
| 46473 | 108 | in Library.foldl (op APPEND) (no_tac, ptacs) end); | 
| 21425 | 109 | |
| 27229 | 110 | fun ptac ctxt prog = let | 
| 32952 | 111 | val proga = maps (atomizeD ctxt) prog (* atomize the prog *) | 
| 21425 | 112 | in (REPEAT_DETERM1 o FIRST' [ | 
| 60754 | 113 | resolve_tac ctxt [TrueI], (* "True" *) | 
| 114 | resolve_tac ctxt [conjI], (* "[| P; Q |] ==> P & Q" *) | |
| 115 | resolve_tac ctxt [allI], (* "(!!x. P x) ==> ! x. P x" *) | |
| 116 | resolve_tac ctxt [exI], (* "P x ==> ? x. P x" *) | |
| 117 | resolve_tac ctxt [impI] THEN' (* "(P ==> Q) ==> P --> Q" *) | |
| 52088 | 118 | asm_full_simp_tac (put_simpset atomize_ss ctxt) THEN' (* atomize the asms *) | 
| 60754 | 119 | (REPEAT_DETERM o (eresolve_tac ctxt [conjE])) (* split the asms *) | 
| 21425 | 120 | ]) | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
56467diff
changeset | 121 | ORELSE' resolve_tac ctxt [disjI1,disjI2] (* "P ==> P | Q","Q ==> P | Q"*) | 
| 60754 | 122 | ORELSE' ((resolve_tac ctxt proga APPEND' hyp_resolve_tac ctxt) | 
| 21425 | 123 | THEN' (fn _ => check_HOHH_tac2)) | 
| 124 | end; | |
| 125 | ||
| 27229 | 126 | fun prolog_tac ctxt prog = | 
| 127 | check_HOHH_tac1 THEN | |
| 128 | DEPTH_SOLVE (ptac ctxt (map check_HOHH prog) 1); | |
| 21425 | 129 | |
| 130 | val prog_HOHH = []; | |
| 131 | ||
| 132 | end; |