src/HOL/Tools/reconstruction.ML
author berghofe
Thu, 21 Apr 2005 18:57:18 +0200
changeset 15794 5de27a5fc5ed
parent 15684 5ec4d21889d6
child 15955 87cf2ce8ede8
permissions -rw-r--r--
Adapted to new interface of instantiation and unification / matching functions.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
15151
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Reconstruction.thy
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
     2
    ID: $Id$
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson and Claire Quigley
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
     4
    Copyright   2004  University of Cambridge
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
     5
*)
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
     6
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
     7
(*Attributes for reconstructing external resolution proofs*)
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
     8
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
     9
structure Reconstruction =
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    10
let open Attrib
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    11
in
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    12
struct
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    13
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    14
(**************************************************************)
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    15
(* extra functions necessary for factoring and paramodulation *)
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    16
(**************************************************************)
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    17
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    18
fun mksubstlist [] sublist = sublist
15794
5de27a5fc5ed Adapted to new interface of instantiation and unification / matching functions.
berghofe
parents: 15684
diff changeset
    19
  | mksubstlist ((a, (_, b)) :: rest) sublist = 
15151
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    20
      let val vartype = type_of b
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    21
          val avar = Var(a,vartype)
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    22
          val newlist = ((avar,b)::sublist) 
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    23
      in mksubstlist rest newlist end;
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    24
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    25
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    26
(**** attributes ****)
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    27
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    28
(** Binary resolution **)
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    29
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    30
fun binary_rule ((cl1, lit1), (cl2 , lit2)) =
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    31
     select_literal (lit1 + 1) cl1
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    32
     RSN ((lit2 + 1), cl2);
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    33
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    34
fun binary_syntax ((i, B), j) (x, A) = (x, binary_rule ((A,i), (B,j)));
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    35
15384
b13eb8a8897d renamed attributes to lower case
paulson
parents: 15365
diff changeset
    36
fun gen_binary thm = syntax
15151
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    37
      ((Scan.lift Args.nat -- thm -- Scan.lift Args.nat) >> binary_syntax);
15384
b13eb8a8897d renamed attributes to lower case
paulson
parents: 15365
diff changeset
    38
val binary_global = gen_binary global_thm;
b13eb8a8897d renamed attributes to lower case
paulson
parents: 15365
diff changeset
    39
val binary_local = gen_binary local_thm;
15151
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    40
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    41
(*I have not done the MRR rule because it seems to be identifical to 
15384
b13eb8a8897d renamed attributes to lower case
paulson
parents: 15365
diff changeset
    42
binary*)
15151
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    43
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    44
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    45
fun inst_single sign t1 t2 cl =
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    46
    let val ct1 = cterm_of sign t1 and ct2 = cterm_of sign t2
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    47
    in  hd (Seq.list_of(distinct_subgoals_tac
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    48
			    (cterm_instantiate [(ct1,ct2)] cl)))  
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    49
    end;
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    50
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    51
fun inst_subst sign substs cl =
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    52
    if (is_Var (fst(hd(substs))))
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    53
    then inst_single sign (fst (hd substs)) (snd (hd substs)) cl
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    54
    else if (is_Var (snd(hd(substs))))
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    55
    then inst_single sign (snd (hd substs)) (fst (hd substs)) cl
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    56
    else raise THM ("inst_subst", 0, [cl]);
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    57
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    58
(*Grabs the environment from the result of Unify.unifiers*)
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    59
fun getnewenv thisseq = fst (hd (Seq.list_of thisseq));
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    60
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    61
(** Factoring **)
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    62
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    63
fun factor_rule (cl, lit1, lit2) =
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    64
    let
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    65
       val prems = prems_of cl
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    66
       val fac1 = List.nth (prems,lit1)
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    67
       val fac2 = List.nth (prems,lit2)
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    68
       val sign = sign_of_thm cl
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    69
       val unif_env = Unify.unifiers (sign, Envir.empty 0, [(fac1, fac2)])
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    70
       val newenv = getnewenv unif_env
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    71
       val envlist = Envir.alist_of newenv
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    72
     in
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    73
       inst_subst sign (mksubstlist envlist []) cl
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    74
    end;
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    75
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    76
fun factor_syntax (i, j) (x, A) = (x, factor_rule (A,i,j));
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    77
15384
b13eb8a8897d renamed attributes to lower case
paulson
parents: 15365
diff changeset
    78
fun factor x = syntax ((Scan.lift (Args.nat -- Args.nat)) >> factor_syntax) x;
15151
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    79
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    80
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    81
(** Paramodulation **)
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    82
15449
a27c81bd838d fixed the treatment of demodulation and paramodulation
paulson
parents: 15384
diff changeset
    83
(*subst with premises exchanged: that way, side literals of the equality will appear
a27c81bd838d fixed the treatment of demodulation and paramodulation
paulson
parents: 15384
diff changeset
    84
  as the second to last premises of the result.*)
a27c81bd838d fixed the treatment of demodulation and paramodulation
paulson
parents: 15384
diff changeset
    85
val rev_subst = rotate_prems 1 subst;
a27c81bd838d fixed the treatment of demodulation and paramodulation
paulson
parents: 15384
diff changeset
    86
15499
419dc5ffe8bc clausification and proof reconstruction
paulson
parents: 15495
diff changeset
    87
fun paramod_rule ((cl1, lit1), (cl2, lit2)) =
15449
a27c81bd838d fixed the treatment of demodulation and paramodulation
paulson
parents: 15384
diff changeset
    88
    let  val eq_lit_th = select_literal (lit1+1) cl1
a27c81bd838d fixed the treatment of demodulation and paramodulation
paulson
parents: 15384
diff changeset
    89
         val mod_lit_th = select_literal (lit2+1) cl2
a27c81bd838d fixed the treatment of demodulation and paramodulation
paulson
parents: 15384
diff changeset
    90
         val eqsubst = eq_lit_th RSN (2,rev_subst)
a27c81bd838d fixed the treatment of demodulation and paramodulation
paulson
parents: 15384
diff changeset
    91
         val newth = Seq.hd (biresolution false [(false, mod_lit_th)] 1 eqsubst)
15499
419dc5ffe8bc clausification and proof reconstruction
paulson
parents: 15495
diff changeset
    92
         val newth' = Seq.hd (flexflex_rule newth)
15579
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15531
diff changeset
    93
    in Meson.negated_asm_of_head newth' end;
15151
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    94
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    95
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    96
fun paramod_syntax ((i, B), j) (x, A) = (x, paramod_rule ((A,i), (B,j)));
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    97
15384
b13eb8a8897d renamed attributes to lower case
paulson
parents: 15365
diff changeset
    98
fun gen_paramod thm = syntax
15151
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
    99
      ((Scan.lift Args.nat -- thm -- Scan.lift Args.nat) >> paramod_syntax);
15384
b13eb8a8897d renamed attributes to lower case
paulson
parents: 15365
diff changeset
   100
val paramod_global = gen_paramod global_thm;
b13eb8a8897d renamed attributes to lower case
paulson
parents: 15365
diff changeset
   101
val paramod_local = gen_paramod local_thm;
15151
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
   102
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
   103
15449
a27c81bd838d fixed the treatment of demodulation and paramodulation
paulson
parents: 15384
diff changeset
   104
(** Demodulation: rewriting of a single literal (Non-Unit Rewriting, SPASS) **)
15151
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
   105
15449
a27c81bd838d fixed the treatment of demodulation and paramodulation
paulson
parents: 15384
diff changeset
   106
fun demod_rule ((cl1, lit1), (cl2 , lit2)) = 
15151
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
   107
    let  val eq_lit_th = select_literal (lit1+1) cl1
15449
a27c81bd838d fixed the treatment of demodulation and paramodulation
paulson
parents: 15384
diff changeset
   108
         val mod_lit_th = select_literal (lit2+1) cl2
15495
50fde6f04e4c new treatment of demodulation in proof reconstruction
paulson
parents: 15466
diff changeset
   109
	 val (fmod_th,thaw) = Drule.freeze_thaw_robust mod_lit_th
15449
a27c81bd838d fixed the treatment of demodulation and paramodulation
paulson
parents: 15384
diff changeset
   110
         val eqsubst = eq_lit_th RSN (2,rev_subst)
15495
50fde6f04e4c new treatment of demodulation in proof reconstruction
paulson
parents: 15466
diff changeset
   111
         val newth = Seq.hd(biresolution false [(false, fmod_th)] 1 eqsubst)
15499
419dc5ffe8bc clausification and proof reconstruction
paulson
parents: 15495
diff changeset
   112
         val offset = #maxidx(rep_thm newth) + 1
419dc5ffe8bc clausification and proof reconstruction
paulson
parents: 15495
diff changeset
   113
	                  (*ensures "renaming apart" even when Vars are frozen*)
15579
32bee18c675f Tools/meson.ML: signature, structure and "open" rather than "local"
paulson
parents: 15531
diff changeset
   114
    in Meson.negated_asm_of_head (thaw offset newth) end;
15151
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
   115
15449
a27c81bd838d fixed the treatment of demodulation and paramodulation
paulson
parents: 15384
diff changeset
   116
fun demod_syntax ((i, B), j) (x, A) = (x, demod_rule ((A,i), (B,j)));
15151
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
   117
15449
a27c81bd838d fixed the treatment of demodulation and paramodulation
paulson
parents: 15384
diff changeset
   118
fun gen_demod thm = syntax
a27c81bd838d fixed the treatment of demodulation and paramodulation
paulson
parents: 15384
diff changeset
   119
      ((Scan.lift Args.nat -- thm -- Scan.lift Args.nat) >> demod_syntax);
15384
b13eb8a8897d renamed attributes to lower case
paulson
parents: 15365
diff changeset
   120
val demod_global = gen_demod global_thm;
b13eb8a8897d renamed attributes to lower case
paulson
parents: 15365
diff changeset
   121
val demod_local = gen_demod local_thm;
15151
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
   122
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
   123
15359
8bad1f42fec0 new CLAUSIFY attribute for proof reconstruction with lemmas
paulson
parents: 15151
diff changeset
   124
(** Conversion of a theorem into clauses **)
8bad1f42fec0 new CLAUSIFY attribute for proof reconstruction with lemmas
paulson
parents: 15151
diff changeset
   125
15466
dce7827f8d75 implemented cache for conversion to clauses
paulson
parents: 15449
diff changeset
   126
local
dce7827f8d75 implemented cache for conversion to clauses
paulson
parents: 15449
diff changeset
   127
dce7827f8d75 implemented cache for conversion to clauses
paulson
parents: 15449
diff changeset
   128
  (*Cache for clauses: could be a hash table if we provided them.*)
dce7827f8d75 implemented cache for conversion to clauses
paulson
parents: 15449
diff changeset
   129
  val cc = ref (Symtab.empty : (thm * thm list) Symtab.table)
dce7827f8d75 implemented cache for conversion to clauses
paulson
parents: 15449
diff changeset
   130
dce7827f8d75 implemented cache for conversion to clauses
paulson
parents: 15449
diff changeset
   131
  fun memo_cnf th =
dce7827f8d75 implemented cache for conversion to clauses
paulson
parents: 15449
diff changeset
   132
    case Thm.name_of_thm th of
15499
419dc5ffe8bc clausification and proof reconstruction
paulson
parents: 15495
diff changeset
   133
        "" => ResAxioms.meta_cnf_axiom th (*no name, so can't cache*)
15466
dce7827f8d75 implemented cache for conversion to clauses
paulson
parents: 15449
diff changeset
   134
      | s  => case Symtab.lookup (!cc,s) of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15499
diff changeset
   135
      	        NONE => 
15499
419dc5ffe8bc clausification and proof reconstruction
paulson
parents: 15495
diff changeset
   136
      	        	let val cls = ResAxioms.meta_cnf_axiom th
15466
dce7827f8d75 implemented cache for conversion to clauses
paulson
parents: 15449
diff changeset
   137
      	        	in  cc := Symtab.update ((s, (th,cls)), !cc); cls
dce7827f8d75 implemented cache for conversion to clauses
paulson
parents: 15449
diff changeset
   138
      	        	end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15499
diff changeset
   139
      	      | SOME(th',cls) =>
15466
dce7827f8d75 implemented cache for conversion to clauses
paulson
parents: 15449
diff changeset
   140
      	      	    if eq_thm(th,th') then cls
dce7827f8d75 implemented cache for conversion to clauses
paulson
parents: 15449
diff changeset
   141
      	      	    else (*New theorem stored under the same name? Possible??*)
15499
419dc5ffe8bc clausification and proof reconstruction
paulson
parents: 15495
diff changeset
   142
      	      	      let val cls = ResAxioms.meta_cnf_axiom th
15466
dce7827f8d75 implemented cache for conversion to clauses
paulson
parents: 15449
diff changeset
   143
      	        	  in  cc := Symtab.update ((s, (th,cls)), !cc); cls
dce7827f8d75 implemented cache for conversion to clauses
paulson
parents: 15449
diff changeset
   144
      	        	  end;
dce7827f8d75 implemented cache for conversion to clauses
paulson
parents: 15449
diff changeset
   145
dce7827f8d75 implemented cache for conversion to clauses
paulson
parents: 15449
diff changeset
   146
in
15499
419dc5ffe8bc clausification and proof reconstruction
paulson
parents: 15495
diff changeset
   147
fun clausify_rule (A,i) = List.nth (memo_cnf A,i)
15466
dce7827f8d75 implemented cache for conversion to clauses
paulson
parents: 15449
diff changeset
   148
end;
15359
8bad1f42fec0 new CLAUSIFY attribute for proof reconstruction with lemmas
paulson
parents: 15151
diff changeset
   149
8bad1f42fec0 new CLAUSIFY attribute for proof reconstruction with lemmas
paulson
parents: 15151
diff changeset
   150
fun clausify_syntax i (x, A) = (x, clausify_rule (A,i));
8bad1f42fec0 new CLAUSIFY attribute for proof reconstruction with lemmas
paulson
parents: 15151
diff changeset
   151
15384
b13eb8a8897d renamed attributes to lower case
paulson
parents: 15365
diff changeset
   152
fun clausify x = syntax ((Scan.lift Args.nat) >> clausify_syntax) x;
15359
8bad1f42fec0 new CLAUSIFY attribute for proof reconstruction with lemmas
paulson
parents: 15151
diff changeset
   153
8bad1f42fec0 new CLAUSIFY attribute for proof reconstruction with lemmas
paulson
parents: 15151
diff changeset
   154
15151
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
   155
(** theory setup **)
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
   156
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
   157
val setup =
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
   158
  [Attrib.add_attributes
15384
b13eb8a8897d renamed attributes to lower case
paulson
parents: 15365
diff changeset
   159
     [("binary", (binary_global, binary_local), "binary resolution"),
b13eb8a8897d renamed attributes to lower case
paulson
parents: 15365
diff changeset
   160
      ("paramod", (paramod_global, paramod_local), "paramodulation"),
b13eb8a8897d renamed attributes to lower case
paulson
parents: 15365
diff changeset
   161
      ("demod", (demod_global, demod_local), "demodulation"),
b13eb8a8897d renamed attributes to lower case
paulson
parents: 15365
diff changeset
   162
      ("factor", (factor, factor), "factoring"),
b13eb8a8897d renamed attributes to lower case
paulson
parents: 15365
diff changeset
   163
      ("clausify", (clausify, clausify), "conversion to clauses")]];
15151
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
   164
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
   165
end
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
   166
end
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
   167
429666b09783 proof reconstruction for external ATPs
paulson
parents:
diff changeset
   168