src/HOL/Prolog/prolog.ML
author paulson <lp15@cam.ac.uk>
Thu, 22 Feb 2018 22:58:27 +0000
changeset 67689 2c38ffd6ec71
parent 60754 02924903a6fd
child 69597 ff784d5a5bfb
permissions -rw-r--r--
type class linordered_nonzero_semiring has new axiom to guarantee characteristic 0
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
     1
(*  Title:    HOL/Prolog/prolog.ML
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
     2
    Author:   David von Oheimb (based on a lecture on Lambda Prolog by Nadathur)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
     3
*)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
     4
56467
8d7d6f17c6a7 more uniform ML/document antiquotations;
wenzelm
parents: 56245
diff changeset
     5
Options.default_put_bool @{system_option show_main_goal} true;
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
     6
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
     7
structure Prolog =
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
     8
struct
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
     9
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    10
exception not_HOHH;
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    11
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    12
fun isD t = case t of
38557
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
    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
56245
84fc7dfa3cd4 more qualified names;
wenzelm
parents: 55151
diff changeset
    16
  | Const(@{const_name Pure.imp},_)$l$r     => isG l andalso isD r
38557
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
    17
  | Const(@{const_name All},_)$Abs(s,_,t) => isD t
56245
84fc7dfa3cd4 more qualified names;
wenzelm
parents: 55151
diff changeset
    18
  | Const(@{const_name Pure.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
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
    20
  | Const(@{const_name Ex} ,_)$_          => false
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
    21
  | Const(@{const_name Not},_)$_          => false
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
    22
  | Const(@{const_name True},_)           => false
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
    23
  | Const(@{const_name False},_)          => false
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    24
  | l $ r                     => isD l
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    25
  | Const _ (* rigid atom *)  => true
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    26
  | Bound _ (* rigid atom *)  => true
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    27
  | Free  _ (* rigid atom *)  => true
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    28
  | _    (* flexible atom,
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    29
            anything else *)  => false
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    30
and
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    31
    isG t = case t of
38557
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
    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
56245
84fc7dfa3cd4 more qualified names;
wenzelm
parents: 55151
diff changeset
    36
  | Const(@{const_name Pure.imp},_)$l$r     => isD l andalso isG r
38557
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
    37
  | Const(@{const_name All},_)$Abs(_,_,t) => isG t
56245
84fc7dfa3cd4 more qualified names;
wenzelm
parents: 55151
diff changeset
    38
  | Const(@{const_name Pure.all},_)$Abs(_,_,t) => isG t
38557
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
    39
  | Const(@{const_name Ex} ,_)$Abs(_,_,t) => isG t
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
    40
  | Const(@{const_name True},_)           => true
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
    41
  | Const(@{const_name Not},_)$_          => false
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
    42
  | Const(@{const_name False},_)          => false
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    43
  | _ (* atom *)              => true;
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    44
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    45
val check_HOHH_tac1 = PRIMITIVE (fn thm =>
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
    46
        if isG (Thm.concl_of thm) then thm else raise not_HOHH);
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    47
val check_HOHH_tac2 = PRIMITIVE (fn thm =>
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
    48
        if forall isG (Thm.prems_of thm) then thm else raise not_HOHH);
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
    49
fun check_HOHH thm  = (if isD (Thm.concl_of thm) andalso forall isG (Thm.prems_of thm)
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    50
                        then thm else raise not_HOHH);
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    51
55143
04448228381d explicit eigen-context for attributes "where", "of", and corresponding read_instantiate, instantiate_tac;
wenzelm
parents: 52233
diff changeset
    52
fun atomizeD ctxt thm =
04448228381d explicit eigen-context for attributes "where", "of", and corresponding read_instantiate, instantiate_tac;
wenzelm
parents: 52233
diff changeset
    53
  let
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
    54
    fun at  thm = case Thm.concl_of thm of
55143
04448228381d explicit eigen-context for attributes "where", "of", and corresponding read_instantiate, instantiate_tac;
wenzelm
parents: 52233
diff changeset
    55
      _$(Const(@{const_name All} ,_)$Abs(s,_,_))=>
59755
f8d164ab0dc1 more position information;
wenzelm
parents: 59582
diff changeset
    56
        let val s' = if s="P" then "PP" else s in
f8d164ab0dc1 more position information;
wenzelm
parents: 59582
diff changeset
    57
          at(thm RS (Rule_Insts.read_instantiate ctxt [((("x", 0), Position.none), s')] [s'] spec))
f8d164ab0dc1 more position information;
wenzelm
parents: 59582
diff changeset
    58
        end
38795
848be46708dc formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents: 38786
diff changeset
    59
    | _$(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
    60
    | _$(Const(@{const_name HOL.implies},_)$_$_)     => at(thm RS mp)
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    61
    | _                             => [thm]
55143
04448228381d explicit eigen-context for attributes "where", "of", and corresponding read_instantiate, instantiate_tac;
wenzelm
parents: 52233
diff changeset
    62
  in map zero_var_indexes (at thm) end;
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    63
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    64
val atomize_ss =
52233
wenzelm
parents: 52088
diff changeset
    65
  (empty_simpset @{context} |> Simplifier.set_mksimps (mksimps mksimps_pairs))
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    66
  addsimps [
45654
cf10bde35973 more antiquotations;
wenzelm
parents: 45625
diff changeset
    67
        @{thm all_conj_distrib}, (* "(! x. P x & Q x) = ((! x. P x) & (! x. Q x))" *)
46161
4ed94d92ae19 prefer antiquotations;
wenzelm
parents: 45654
diff changeset
    68
        @{thm imp_conjL} RS sym, (* "(D :- G1 :- G2) = (D :- G1 & G2)" *)
4ed94d92ae19 prefer antiquotations;
wenzelm
parents: 45654
diff changeset
    69
        @{thm imp_conjR},        (* "(D1 & D2 :- G) = ((D1 :- G) & (D2 :- G))" *)
52088
7d8b53e80ce7 proper run-time context;
wenzelm
parents: 52043
diff changeset
    70
        @{thm imp_all}]          (* "((!x. D) :- G) = (!x. D :- G)" *)
7d8b53e80ce7 proper run-time context;
wenzelm
parents: 52043
diff changeset
    71
  |> simpset_of;
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    72
45625
750c5a47400b modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents: 39125
diff changeset
    73
32283
3bebc195c124 qualified Subgoal.FOCUS;
wenzelm
parents: 32282
diff changeset
    74
(*val hyp_resolve_tac = Subgoal.FOCUS_PREMS (fn {prems, ...} =>
32260
eb97888fa422 eliminated METAHYPS;
wenzelm
parents: 32010
diff changeset
    75
                                  resolve_tac (maps atomizeD prems) 1);
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    76
  -- is nice, but cannot instantiate unknowns in the assumptions *)
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59755
diff changeset
    77
fun hyp_resolve_tac ctxt = SUBGOAL (fn (subgoal, i) =>
46473
a687b75f9fa8 more conventional tactic setup;
wenzelm
parents: 46161
diff changeset
    78
  let
38557
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
    79
        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
    80
        |   ap (Const(@{const_name HOL.implies},_)$_$t)    =(case ap t of (k,_,t) => (k,true ,t))
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    81
        |   ap t                          =                         (0,false,t);
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    82
(*
56245
84fc7dfa3cd4 more qualified names;
wenzelm
parents: 55151
diff changeset
    83
        fun rep_goal (Const (@{const_name Pure.all},_)$Abs (_,_,t)) = rep_goal t
84fc7dfa3cd4 more qualified names;
wenzelm
parents: 55151
diff changeset
    84
        |   rep_goal (Const (@{const_name Pure.imp},_)$s$t)         =
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    85
                        (case rep_goal t of (l,t) => (s::l,t))
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    86
        |   rep_goal t                             = ([]  ,t);
38557
9926c47ad1a1 more antiquotations
haftmann
parents: 38549
diff changeset
    87
        val (prems, Const(@{const_name Trueprop}, _)$concl) = rep_goal
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    88
                                                (#3(dest_state (st,i)));
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    89
*)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    90
        val prems = Logic.strip_assums_hyp subgoal;
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    91
        val concl = HOLogic.dest_Trueprop (Logic.strip_assums_concl subgoal);
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    92
        fun drot_tac k i = DETERM (rotate_tac k i);
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    93
        fun spec_tac 0 i = all_tac
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59755
diff changeset
    94
        |   spec_tac k i = EVERY' [dresolve_tac ctxt [spec], drot_tac ~1, spec_tac (k-1)] i;
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    95
        fun dup_spec_tac k i = if k = 0 then all_tac else EVERY'
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59755
diff changeset
    96
                      [DETERM o (eresolve_tac ctxt [all_dupE]), drot_tac ~2, spec_tac (k-1)] i;
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    97
        fun same_head _ (Const (x,_)) (Const (y,_)) = x = y
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    98
        |   same_head k (s$_)         (t$_)         = same_head k s t
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
    99
        |   same_head k (Bound i)     (Bound j)     = i = j + k
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
   100
        |   same_head _ _             _             = true;
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
   101
        fun mapn f n []      = []
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
   102
        |   mapn f n (x::xs) = f n x::mapn f (n+1) xs;
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
   103
        fun pres_tac (k,arrow,t) n i = drot_tac n i THEN (
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59755
diff changeset
   104
          if same_head k t concl
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59755
diff changeset
   105
          then dup_spec_tac k i THEN
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59755
diff changeset
   106
               (if arrow then eresolve_tac ctxt [mp] i THEN drot_tac (~n) i else assume_tac ctxt i)
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59755
diff changeset
   107
          else no_tac);
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
   108
        val ptacs = mapn (fn n => fn t =>
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
   109
                          pres_tac (ap (HOLogic.dest_Trueprop t)) n i) 0 prems;
46473
a687b75f9fa8 more conventional tactic setup;
wenzelm
parents: 46161
diff changeset
   110
  in Library.foldl (op APPEND) (no_tac, ptacs) end);
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
   111
27229
f656a12e0f4e ptac/prolog_tac: proper context;
wenzelm
parents: 27153
diff changeset
   112
fun ptac ctxt prog = let
32952
aeb1e44fbc19 replaced String.concat by implode;
wenzelm
parents: 32740
diff changeset
   113
  val proga = maps (atomizeD ctxt) prog         (* atomize the prog *)
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
   114
  in    (REPEAT_DETERM1 o FIRST' [
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59755
diff changeset
   115
                resolve_tac ctxt [TrueI],                     (* "True" *)
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59755
diff changeset
   116
                resolve_tac ctxt [conjI],                     (* "[| P; Q |] ==> P & Q" *)
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59755
diff changeset
   117
                resolve_tac ctxt [allI],                      (* "(!!x. P x) ==> ! x. P x" *)
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59755
diff changeset
   118
                resolve_tac ctxt [exI],                       (* "P x ==> ? x. P x" *)
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59755
diff changeset
   119
                resolve_tac ctxt [impI] THEN'                 (* "(P ==> Q) ==> P --> Q" *)
52088
7d8b53e80ce7 proper run-time context;
wenzelm
parents: 52043
diff changeset
   120
                  asm_full_simp_tac (put_simpset atomize_ss ctxt) THEN'    (* atomize the asms *)
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59755
diff changeset
   121
                  (REPEAT_DETERM o (eresolve_tac ctxt [conjE]))        (* split the asms *)
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
   122
                ])
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 56467
diff changeset
   123
        ORELSE' resolve_tac ctxt [disjI1,disjI2]     (* "P ==> P | Q","Q ==> P | Q"*)
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 59755
diff changeset
   124
        ORELSE' ((resolve_tac ctxt proga APPEND' hyp_resolve_tac ctxt)
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
   125
                 THEN' (fn _ => check_HOHH_tac2))
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
   126
end;
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
   127
27229
f656a12e0f4e ptac/prolog_tac: proper context;
wenzelm
parents: 27153
diff changeset
   128
fun prolog_tac ctxt prog =
f656a12e0f4e ptac/prolog_tac: proper context;
wenzelm
parents: 27153
diff changeset
   129
  check_HOHH_tac1 THEN
f656a12e0f4e ptac/prolog_tac: proper context;
wenzelm
parents: 27153
diff changeset
   130
  DEPTH_SOLVE (ptac ctxt (map check_HOHH prog) 1);
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
   131
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
   132
val prog_HOHH = [];
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
   133
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents:
diff changeset
   134
end;