| author | huffman | 
| Tue, 06 Sep 2011 07:48:59 -0700 | |
| changeset 44749 | 5b1e1432c320 | 
| parent 39125 | f45d332a90e3 | 
| child 45625 | 750c5a47400b | 
| 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 | ||
| 39125 
f45d332a90e3
pretty_goals: turned some global references and function arguments into configuration options (goals_limit = 10, goals_total = true, show_main_goal = false) depending on the context;
 wenzelm parents: 
38798diff
changeset | 5 | Goal_Display.show_main_goal_default := true; | 
| 21425 | 6 | |
| 7 | structure Prolog = | |
| 8 | struct | |
| 9 | ||
| 10 | exception not_HOHH; | |
| 11 | ||
| 12 | fun isD t = case t of | |
| 38557 | 13 |     Const(@{const_name Trueprop},_)$t     => isD t
 | 
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 14 |   | Const(@{const_name HOL.conj}  ,_)$l$r     => isD l andalso isD r
 | 
| 38786 
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
 haftmann parents: 
38557diff
changeset | 15 |   | Const(@{const_name HOL.implies},_)$l$r     => isG l andalso isD r
 | 
| 21425 | 16 | | Const( "==>",_)$l$r => isG l andalso isD r | 
| 38557 | 17 |   | Const(@{const_name All},_)$Abs(s,_,t) => isD t
 | 
| 21425 | 18 |   | Const("all",_)$Abs(s,_,t) => isD t
 | 
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 19 |   | Const(@{const_name HOL.disj},_)$_$_       => false
 | 
| 38557 | 20 |   | Const(@{const_name Ex} ,_)$_          => false
 | 
| 21 |   | Const(@{const_name Not},_)$_          => false
 | |
| 22 |   | Const(@{const_name True},_)           => false
 | |
| 23 |   | Const(@{const_name False},_)          => false
 | |
| 21425 | 24 | | l $ r => isD l | 
| 25 | | Const _ (* rigid atom *) => true | |
| 26 | | Bound _ (* rigid atom *) => true | |
| 27 | | Free _ (* rigid atom *) => true | |
| 28 | | _ (* flexible atom, | |
| 29 | anything else *) => false | |
| 30 | and | |
| 31 | isG t = case t of | |
| 38557 | 32 |     Const(@{const_name Trueprop},_)$t     => isG t
 | 
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 33 |   | Const(@{const_name HOL.conj}  ,_)$l$r     => isG l andalso isG r
 | 
| 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 34 |   | Const(@{const_name HOL.disj}  ,_)$l$r     => isG l andalso isG r
 | 
| 38786 
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
 haftmann parents: 
38557diff
changeset | 35 |   | Const(@{const_name HOL.implies},_)$l$r     => isD l andalso isG r
 | 
| 21425 | 36 | | Const( "==>",_)$l$r => isD l andalso isG r | 
| 38557 | 37 |   | Const(@{const_name All},_)$Abs(_,_,t) => isG t
 | 
| 21425 | 38 |   | Const("all",_)$Abs(_,_,t) => isG t
 | 
| 38557 | 39 |   | Const(@{const_name Ex} ,_)$Abs(_,_,t) => isG t
 | 
| 40 |   | Const(@{const_name True},_)           => true
 | |
| 41 |   | Const(@{const_name Not},_)$_          => false
 | |
| 42 |   | Const(@{const_name False},_)          => false
 | |
| 21425 | 43 | | _ (* atom *) => true; | 
| 44 | ||
| 45 | val check_HOHH_tac1 = PRIMITIVE (fn thm => | |
| 46 | if isG (concl_of thm) then thm else raise not_HOHH); | |
| 47 | val check_HOHH_tac2 = PRIMITIVE (fn thm => | |
| 48 | if forall isG (prems_of thm) then thm else raise not_HOHH); | |
| 49 | fun check_HOHH thm = (if isD (concl_of thm) andalso forall isG (prems_of thm) | |
| 50 | then thm else raise not_HOHH); | |
| 51 | ||
| 27229 | 52 | fun atomizeD ctxt thm = let | 
| 21425 | 53 | fun at thm = case concl_of thm of | 
| 38557 | 54 |       _$(Const(@{const_name All} ,_)$Abs(s,_,_))=> at(thm RS
 | 
| 27239 | 55 |         (read_instantiate ctxt [(("x", 0), "?" ^ (if s="P" then "PP" else s))] spec))
 | 
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 56 |     | _$(Const(@{const_name HOL.conj},_)$_$_)       => at(thm RS conjunct1)@at(thm RS conjunct2)
 | 
| 38786 
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
 haftmann parents: 
38557diff
changeset | 57 |     | _$(Const(@{const_name HOL.implies},_)$_$_)     => at(thm RS mp)
 | 
| 21425 | 58 | | _ => [thm] | 
| 59 | in map zero_var_indexes (at thm) end; | |
| 60 | ||
| 61 | val atomize_ss = | |
| 35232 
f588e1169c8b
renamed Simplifier.theory_context to Simplifier.global_context to emphasize that this is not the real thing;
 wenzelm parents: 
32952diff
changeset | 62 |   Simplifier.global_context @{theory} empty_ss
 | 
| 21425 | 63 | setmksimps (mksimps mksimps_pairs) | 
| 64 | addsimps [ | |
| 65 | all_conj_distrib, (* "(! x. P x & Q x) = ((! x. P x) & (! x. Q x))" *) | |
| 66 | imp_conjL RS sym, (* "(D :- G1 :- G2) = (D :- G1 & G2)" *) | |
| 67 | imp_conjR, (* "(D1 & D2 :- G) = ((D1 :- G) & (D2 :- G))" *) | |
| 68 | imp_all]; (* "((!x. D) :- G) = (!x. D :- G)" *) | |
| 69 | ||
| 32283 | 70 | (*val hyp_resolve_tac = Subgoal.FOCUS_PREMS (fn {prems, ...} =>
 | 
| 32260 | 71 | resolve_tac (maps atomizeD prems) 1); | 
| 21425 | 72 | -- is nice, but cannot instantiate unknowns in the assumptions *) | 
| 73 | fun hyp_resolve_tac i st = let | |
| 38557 | 74 |         fun ap (Const(@{const_name All},_)$Abs(_,_,t))=(case ap t of (k,a,t) => (k+1,a  ,t))
 | 
| 38786 
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
 haftmann parents: 
38557diff
changeset | 75 |         |   ap (Const(@{const_name HOL.implies},_)$_$t)    =(case ap t of (k,_,t) => (k,true ,t))
 | 
| 21425 | 76 | | ap t = (0,false,t); | 
| 77 | (* | |
| 78 |         fun rep_goal (Const ("all",_)$Abs (_,_,t)) = rep_goal t
 | |
| 79 |         |   rep_goal (Const ("==>",_)$s$t)         =
 | |
| 80 | (case rep_goal t of (l,t) => (s::l,t)) | |
| 81 | | rep_goal t = ([] ,t); | |
| 38557 | 82 |         val (prems, Const(@{const_name Trueprop}, _)$concl) = rep_goal
 | 
| 21425 | 83 | (#3(dest_state (st,i))); | 
| 84 | *) | |
| 36945 | 85 | val subgoal = #3(Thm.dest_state (st,i)); | 
| 21425 | 86 | val prems = Logic.strip_assums_hyp subgoal; | 
| 87 | val concl = HOLogic.dest_Trueprop (Logic.strip_assums_concl subgoal); | |
| 88 | fun drot_tac k i = DETERM (rotate_tac k i); | |
| 89 | fun spec_tac 0 i = all_tac | |
| 90 | | spec_tac k i = EVERY' [dtac spec, drot_tac ~1, spec_tac (k-1)] i; | |
| 91 | fun dup_spec_tac k i = if k = 0 then all_tac else EVERY' | |
| 92 | [DETERM o (etac all_dupE), drot_tac ~2, spec_tac (k-1)] i; | |
| 93 | fun same_head _ (Const (x,_)) (Const (y,_)) = x = y | |
| 94 | | same_head k (s$_) (t$_) = same_head k s t | |
| 95 | | same_head k (Bound i) (Bound j) = i = j + k | |
| 96 | | same_head _ _ _ = true; | |
| 97 | fun mapn f n [] = [] | |
| 98 | | mapn f n (x::xs) = f n x::mapn f (n+1) xs; | |
| 99 | fun pres_tac (k,arrow,t) n i = drot_tac n i THEN ( | |
| 100 | if same_head k t concl | |
| 101 | then dup_spec_tac k i THEN | |
| 102 | (if arrow then etac mp i THEN drot_tac (~n) i else atac i) | |
| 103 | else no_tac); | |
| 104 | val ptacs = mapn (fn n => fn t => | |
| 105 | pres_tac (ap (HOLogic.dest_Trueprop t)) n i) 0 prems; | |
| 106 | in Library.foldl (op APPEND) (no_tac, ptacs) st end; | |
| 107 | ||
| 27229 | 108 | fun ptac ctxt prog = let | 
| 32952 | 109 | val proga = maps (atomizeD ctxt) prog (* atomize the prog *) | 
| 21425 | 110 | in (REPEAT_DETERM1 o FIRST' [ | 
| 111 | rtac TrueI, (* "True" *) | |
| 112 | rtac conjI, (* "[| P; Q |] ==> P & Q" *) | |
| 113 | rtac allI, (* "(!!x. P x) ==> ! x. P x" *) | |
| 114 | rtac exI, (* "P x ==> ? x. P x" *) | |
| 115 | rtac impI THEN' (* "(P ==> Q) ==> P --> Q" *) | |
| 116 | asm_full_simp_tac atomize_ss THEN' (* atomize the asms *) | |
| 117 | (REPEAT_DETERM o (etac conjE)) (* split the asms *) | |
| 118 | ]) | |
| 119 | ORELSE' resolve_tac [disjI1,disjI2] (* "P ==> P | Q","Q ==> P | Q"*) | |
| 120 | ORELSE' ((resolve_tac proga APPEND' hyp_resolve_tac) | |
| 121 | THEN' (fn _ => check_HOHH_tac2)) | |
| 122 | end; | |
| 123 | ||
| 27229 | 124 | fun prolog_tac ctxt prog = | 
| 125 | check_HOHH_tac1 THEN | |
| 126 | DEPTH_SOLVE (ptac ctxt (map check_HOHH prog) 1); | |
| 21425 | 127 | |
| 128 | val prog_HOHH = []; | |
| 129 | ||
| 130 | end; |