author | blanchet |
Thu, 10 Oct 2013 08:23:57 +0200 | |
changeset 54096 | 8ab8794410cd |
parent 52233 | eb84dab7d4c1 |
child 55143 | 04448228381d |
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 |
||
52043
286629271d65
more system options as context-sensitive config options;
wenzelm
parents:
46473
diff
changeset
|
5 |
Options.default_put_bool @{option show_main_goal} 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:
38786
diff
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:
38557
diff
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:
38786
diff
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:
38786
diff
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:
38786
diff
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:
38557
diff
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:
38786
diff
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:
38557
diff
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 = |
|
52233 | 62 |
(empty_simpset @{context} |> Simplifier.set_mksimps (mksimps mksimps_pairs)) |
21425 | 63 |
addsimps [ |
45654 | 64 |
@{thm all_conj_distrib}, (* "(! x. P x & Q x) = ((! x. P x) & (! x. Q x))" *) |
46161 | 65 |
@{thm imp_conjL} RS sym, (* "(D :- G1 :- G2) = (D :- G1 & G2)" *) |
66 |
@{thm imp_conjR}, (* "(D1 & D2 :- G) = ((D1 :- G) & (D2 :- G))" *) |
|
52088 | 67 |
@{thm imp_all}] (* "((!x. D) :- G) = (!x. D :- G)" *) |
68 |
|> simpset_of; |
|
21425 | 69 |
|
45625
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents:
39125
diff
changeset
|
70 |
|
32283 | 71 |
(*val hyp_resolve_tac = Subgoal.FOCUS_PREMS (fn {prems, ...} => |
32260 | 72 |
resolve_tac (maps atomizeD prems) 1); |
21425 | 73 |
-- is nice, but cannot instantiate unknowns in the assumptions *) |
46473 | 74 |
val hyp_resolve_tac = SUBGOAL (fn (subgoal, i) => |
75 |
let |
|
38557 | 76 |
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:
38557
diff
changeset
|
77 |
| ap (Const(@{const_name HOL.implies},_)$_$t) =(case ap t of (k,_,t) => (k,true ,t)) |
21425 | 78 |
| ap t = (0,false,t); |
79 |
(* |
|
80 |
fun rep_goal (Const ("all",_)$Abs (_,_,t)) = rep_goal t |
|
81 |
| rep_goal (Const ("==>",_)$s$t) = |
|
82 |
(case rep_goal t of (l,t) => (s::l,t)) |
|
83 |
| rep_goal t = ([] ,t); |
|
38557 | 84 |
val (prems, Const(@{const_name Trueprop}, _)$concl) = rep_goal |
21425 | 85 |
(#3(dest_state (st,i))); |
86 |
*) |
|
87 |
val prems = Logic.strip_assums_hyp subgoal; |
|
88 |
val concl = HOLogic.dest_Trueprop (Logic.strip_assums_concl subgoal); |
|
89 |
fun drot_tac k i = DETERM (rotate_tac k i); |
|
90 |
fun spec_tac 0 i = all_tac |
|
91 |
| spec_tac k i = EVERY' [dtac spec, drot_tac ~1, spec_tac (k-1)] i; |
|
92 |
fun dup_spec_tac k i = if k = 0 then all_tac else EVERY' |
|
93 |
[DETERM o (etac all_dupE), drot_tac ~2, spec_tac (k-1)] i; |
|
94 |
fun same_head _ (Const (x,_)) (Const (y,_)) = x = y |
|
95 |
| same_head k (s$_) (t$_) = same_head k s t |
|
96 |
| same_head k (Bound i) (Bound j) = i = j + k |
|
97 |
| same_head _ _ _ = true; |
|
98 |
fun mapn f n [] = [] |
|
99 |
| mapn f n (x::xs) = f n x::mapn f (n+1) xs; |
|
100 |
fun pres_tac (k,arrow,t) n i = drot_tac n i THEN ( |
|
101 |
if same_head k t concl |
|
102 |
then dup_spec_tac k i THEN |
|
103 |
(if arrow then etac mp i THEN drot_tac (~n) i else atac i) |
|
104 |
else no_tac); |
|
105 |
val ptacs = mapn (fn n => fn t => |
|
106 |
pres_tac (ap (HOLogic.dest_Trueprop t)) n i) 0 prems; |
|
46473 | 107 |
in Library.foldl (op APPEND) (no_tac, ptacs) end); |
21425 | 108 |
|
27229 | 109 |
fun ptac ctxt prog = let |
32952 | 110 |
val proga = maps (atomizeD ctxt) prog (* atomize the prog *) |
21425 | 111 |
in (REPEAT_DETERM1 o FIRST' [ |
112 |
rtac TrueI, (* "True" *) |
|
113 |
rtac conjI, (* "[| P; Q |] ==> P & Q" *) |
|
114 |
rtac allI, (* "(!!x. P x) ==> ! x. P x" *) |
|
115 |
rtac exI, (* "P x ==> ? x. P x" *) |
|
116 |
rtac impI THEN' (* "(P ==> Q) ==> P --> Q" *) |
|
52088 | 117 |
asm_full_simp_tac (put_simpset atomize_ss ctxt) THEN' (* atomize the asms *) |
21425 | 118 |
(REPEAT_DETERM o (etac conjE)) (* split the asms *) |
119 |
]) |
|
120 |
ORELSE' resolve_tac [disjI1,disjI2] (* "P ==> P | Q","Q ==> P | Q"*) |
|
121 |
ORELSE' ((resolve_tac proga APPEND' hyp_resolve_tac) |
|
122 |
THEN' (fn _ => check_HOHH_tac2)) |
|
123 |
end; |
|
124 |
||
27229 | 125 |
fun prolog_tac ctxt prog = |
126 |
check_HOHH_tac1 THEN |
|
127 |
DEPTH_SOLVE (ptac ctxt (map check_HOHH prog) 1); |
|
21425 | 128 |
|
129 |
val prog_HOHH = []; |
|
130 |
||
131 |
end; |