src/HOL/Prolog/HOHH.ML
author wenzelm
Tue, 18 Oct 2005 17:59:25 +0200
changeset 17892 62c397c17d18
parent 17311 5b1d47d920ce
permissions -rw-r--r--
Simplifier.theory_context;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13208
965f95a3abd9 added the usual file headers
oheimb
parents: 9015
diff changeset
     1
(*  Title:    HOL/Prolog/HOHH.ML
965f95a3abd9 added the usual file headers
oheimb
parents: 9015
diff changeset
     2
    ID:       $Id$
965f95a3abd9 added the usual file headers
oheimb
parents: 9015
diff changeset
     3
    Author:   David von Oheimb (based on a lecture on Lambda Prolog by Nadathur)
965f95a3abd9 added the usual file headers
oheimb
parents: 9015
diff changeset
     4
*)
965f95a3abd9 added the usual file headers
oheimb
parents: 9015
diff changeset
     5
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
     6
exception not_HOHH;
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
     7
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
     8
fun isD t = case t of
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
     9
    Const("Trueprop",_)$t     => isD t
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    10
  | Const("op &"  ,_)$l$r     => isD l andalso isD r
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    11
  | Const("op -->",_)$l$r     => isG l andalso isD r
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    12
  | Const(   "==>",_)$l$r     => isG l andalso isD r
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    13
  | Const("All",_)$Abs(s,_,t) => isD t
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    14
  | Const("all",_)$Abs(s,_,t) => isD t
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    15
  | Const("op |",_)$_$_       => false
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    16
  | Const("Ex" ,_)$_          => false
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    17
  | Const("not",_)$_          => false
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    18
  | Const("True",_)           => false
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    19
  | Const("False",_)          => false
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    20
  | l $ r                     => isD l
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    21
  | Const _ (* rigid atom *)  => true
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    22
  | Bound _ (* rigid atom *)  => true
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    23
  | Free  _ (* rigid atom *)  => true
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    24
  | _    (* flexible atom,
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    25
            anything else *)  => false
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    26
and
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    27
    isG t = case t of
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    28
    Const("Trueprop",_)$t     => isG t
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    29
  | Const("op &"  ,_)$l$r     => isG l andalso isG r
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    30
  | Const("op |"  ,_)$l$r     => isG l andalso isG r
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    31
  | Const("op -->",_)$l$r     => isD l andalso isG r
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    32
  | Const(   "==>",_)$l$r     => isD l andalso isG r
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    33
  | Const("All",_)$Abs(_,_,t) => isG t
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    34
  | Const("all",_)$Abs(_,_,t) => isG t
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    35
  | Const("Ex" ,_)$Abs(_,_,t) => isG t
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    36
  | Const("True",_)           => true
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    37
  | Const("not",_)$_          => false
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    38
  | Const("False",_)          => false
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    39
  | _ (* atom *)              => true;
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    40
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    41
val check_HOHH_tac1 = PRIMITIVE (fn thm =>
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    42
        if isG (concl_of thm) then thm else raise not_HOHH);
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    43
val check_HOHH_tac2 = PRIMITIVE (fn thm =>
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    44
        if forall isG (prems_of thm) then thm else raise not_HOHH);
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    45
fun check_HOHH thm  = (if isD (concl_of thm) andalso forall isG (prems_of thm)
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    46
                        then thm else raise not_HOHH);
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    47
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    48
fun atomizeD thm = let
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    49
    fun at  thm = case concl_of thm of
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    50
      _$(Const("All" ,_)$Abs(s,_,_))=> at(thm RS (read_instantiate [("x",
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    51
                                        "?"^(if s="P" then "PP" else s))] spec))
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    52
    | _$(Const("op &",_)$_$_)       => at(thm RS conjunct1)@at(thm RS conjunct2)
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    53
    | _$(Const("op -->",_)$_$_)     => at(thm RS mp)
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    54
    | _                             => [thm]
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    55
in map zero_var_indexes (at thm) end;
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    56
17892
62c397c17d18 Simplifier.theory_context;
wenzelm
parents: 17311
diff changeset
    57
val atomize_ss =
62c397c17d18 Simplifier.theory_context;
wenzelm
parents: 17311
diff changeset
    58
  Simplifier.theory_context (the_context ()) empty_ss
62c397c17d18 Simplifier.theory_context;
wenzelm
parents: 17311
diff changeset
    59
  setmksimps (mksimps mksimps_pairs)
62c397c17d18 Simplifier.theory_context;
wenzelm
parents: 17311
diff changeset
    60
  addsimps [
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    61
        all_conj_distrib, (* "(! x. P x & Q x) = ((! x. P x) & (! x. Q x))" *)
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    62
        imp_conjL RS sym, (* "(D :- G1 :- G2) = (D :- G1 & G2)" *)
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    63
        imp_conjR,        (* "(D1 & D2 :- G) = ((D1 :- G) & (D2 :- G))" *)
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    64
        imp_all];         (* "((!x. D) :- G) = (!x. D :- G)" *)
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    65
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    66
(*val hyp_resolve_tac = METAHYPS (fn prems =>
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    67
                                  resolve_tac (List.concat (map atomizeD prems)) 1);
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    68
  -- is nice, but cannot instantiate unknowns in the assumptions *)
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    69
fun hyp_resolve_tac i st = let
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    70
        fun ap (Const("All",_)$Abs(_,_,t))=(case ap t of (k,a,t) => (k+1,a  ,t))
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    71
        |   ap (Const("op -->",_)$_$t)    =(case ap t of (k,_,t) => (k,true ,t))
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    72
        |   ap t                          =                         (0,false,t);
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    73
(*
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    74
        fun rep_goal (Const ("all",_)$Abs (_,_,t)) = rep_goal t
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    75
        |   rep_goal (Const ("==>",_)$s$t)         =
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    76
                        (case rep_goal t of (l,t) => (s::l,t))
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    77
        |   rep_goal t                             = ([]  ,t);
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    78
        val (prems, Const("Trueprop", _)$concl) = rep_goal
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    79
                                                (#3(dest_state (st,i)));
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    80
*)
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    81
        val subgoal = #3(dest_state (st,i));
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    82
        val prems = Logic.strip_assums_hyp subgoal;
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    83
        val concl = HOLogic.dest_Trueprop (Logic.strip_assums_concl subgoal);
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    84
        fun drot_tac k i = DETERM (rotate_tac k i);
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    85
        fun spec_tac 0 i = all_tac
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    86
        |   spec_tac k i = EVERY' [dtac spec, drot_tac ~1, spec_tac (k-1)] i;
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    87
        fun dup_spec_tac k i = if k = 0 then all_tac else EVERY'
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    88
                      [DETERM o (etac all_dupE), drot_tac ~2, spec_tac (k-1)] i;
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    89
        fun same_head _ (Const (x,_)) (Const (y,_)) = x = y
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    90
        |   same_head k (s$_)         (t$_)         = same_head k s t
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    91
        |   same_head k (Bound i)     (Bound j)     = i = j + k
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    92
        |   same_head _ _             _             = true;
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    93
        fun mapn f n []      = []
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    94
        |   mapn f n (x::xs) = f n x::mapn f (n+1) xs;
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    95
        fun pres_tac (k,arrow,t) n i = drot_tac n i THEN (
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    96
                if same_head k t concl
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    97
                then dup_spec_tac k i THEN
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    98
                     (if arrow then etac mp i THEN drot_tac (~n) i else atac i)
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
    99
                else no_tac);
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
   100
        val ptacs = mapn (fn n => fn t =>
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
   101
                          pres_tac (ap (HOLogic.dest_Trueprop t)) n i) 0 prems;
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
   102
        in Library.foldl (op APPEND) (no_tac, ptacs) st end;
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
   103
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
   104
fun ptac prog = let
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
   105
  val proga = List.concat (map atomizeD prog)                   (* atomize the prog *)
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
   106
  in    (REPEAT_DETERM1 o FIRST' [
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
   107
                rtac TrueI,                     (* "True" *)
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
   108
                rtac conjI,                     (* "[| P; Q |] ==> P & Q" *)
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
   109
                rtac allI,                      (* "(!!x. P x) ==> ! x. P x" *)
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
   110
                rtac exI,                       (* "P x ==> ? x. P x" *)
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
   111
                rtac impI THEN'                 (* "(P ==> Q) ==> P --> Q" *)
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
   112
                  asm_full_simp_tac atomize_ss THEN'    (* atomize the asms *)
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
   113
                  (REPEAT_DETERM o (etac conjE))        (* split the asms *)
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
   114
                ])
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
   115
        ORELSE' resolve_tac [disjI1,disjI2]     (* "P ==> P | Q","Q ==> P | Q"*)
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
   116
        ORELSE' ((resolve_tac proga APPEND' hyp_resolve_tac)
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
   117
                 THEN' (fn _ => check_HOHH_tac2))
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
   118
end;
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
   119
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
   120
fun prolog_tac prog = check_HOHH_tac1 THEN
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 15570
diff changeset
   121
                      DEPTH_SOLVE (ptac (map check_HOHH prog) 1);
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
   122
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
   123
val prog_HOHH = [];