src/HOL/Tools/res_axioms.ML
author haftmann
Mon, 06 Jun 2005 15:09:47 +0200
changeset 16303 fee0a02f61bb
parent 16173 9e2f6c0a779d
child 16563 a92f96951355
permissions -rw-r--r--
integrated MacOS X installation instructions
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
     1
(*  Author: Jia Meng, Cambridge University Computer Laboratory
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
     2
    ID: $Id$
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
     3
    Copyright 2004 University of Cambridge
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
     4
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
     5
Transformation of axiom rules (elim/intro/etc) into CNF forms.    
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
     6
*)
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
     7
15997
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
     8
signature RES_AXIOMS =
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
     9
  sig
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
    10
  exception ELIMR2FOL of string
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
    11
  val elimRule_tac : thm -> Tactical.tactic
16012
4ae42d8f2fea bug fixes for clause form transformation
paulson
parents: 16009
diff changeset
    12
  val elimR2Fol : thm -> term
15997
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
    13
  val transform_elim : thm -> thm
16039
dfe264950511 Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents: 16012
diff changeset
    14
  val clausify_axiom_pairs : (string*thm) -> (ResClause.clause*thm) list
15997
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
    15
  val cnf_axiom : (string * thm) -> thm list
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
    16
  val meta_cnf_axiom : thm -> thm list
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
    17
  val cnf_rule : thm -> thm list
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
    18
  val cnf_classical_rules_thy : theory -> thm list list * thm list
16039
dfe264950511 Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents: 16012
diff changeset
    19
 
15997
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
    20
  val cnf_simpset_rules_thy : theory -> thm list list * thm list
16039
dfe264950511 Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents: 16012
diff changeset
    21
 
16012
4ae42d8f2fea bug fixes for clause form transformation
paulson
parents: 16009
diff changeset
    22
  val rm_Eps : (term * term) list -> thm list -> term list
15997
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
    23
  val claset_rules_of_thy : theory -> (string * thm) list
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
    24
  val simpset_rules_of_thy : theory -> (string * thm) list
16039
dfe264950511 Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents: 16012
diff changeset
    25
  val clausify_rules_pairs : (string * thm) list -> thm list -> (ResClause.clause * thm) list list * thm list
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
    26
  val setup : (theory -> theory) list
15997
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
    27
  end;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    28
15997
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
    29
structure ResAxioms : RES_AXIOMS =
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
    30
 
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
    31
struct
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    32
15997
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
    33
(**** Transformation of Elimination Rules into First-Order Formulas****)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    34
15390
87f78411f7c9 Comments and other tweaks by Jia
paulson
parents: 15371
diff changeset
    35
(* a tactic used to prove an elim-rule. *)
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
    36
fun elimRule_tac th =
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
    37
    ((rtac impI 1) ORELSE (rtac notI 1)) THEN (etac th 1) THEN
15371
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    38
    REPEAT(Fast_tac 1);
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    39
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    40
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    41
(* This following version fails sometimes, need to investigate, do not use it now. *)
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
    42
fun elimRule_tac' th =
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
    43
   ((rtac impI 1) ORELSE (rtac notI 1)) THEN (etac th 1) THEN
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    44
   REPEAT(SOLVE((etac exI 1) ORELSE (rtac conjI 1) ORELSE (rtac disjI1 1) ORELSE (rtac disjI2 1))); 
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    45
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    46
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    47
exception ELIMR2FOL of string;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    48
15390
87f78411f7c9 Comments and other tweaks by Jia
paulson
parents: 15371
diff changeset
    49
(* functions used to construct a formula *)
87f78411f7c9 Comments and other tweaks by Jia
paulson
parents: 15371
diff changeset
    50
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    51
fun make_disjs [x] = x
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    52
  | make_disjs (x :: xs) = HOLogic.mk_disj(x, make_disjs xs)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    53
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    54
fun make_conjs [x] = x
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    55
  | make_conjs (x :: xs) =  HOLogic.mk_conj(x, make_conjs xs)
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    56
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    57
fun add_EX tm [] = tm
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    58
  | add_EX tm ((x,xtp)::xs) = add_EX (HOLogic.exists_const xtp $ Abs(x,xtp,tm)) xs;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    59
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    60
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    61
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    62
fun is_neg (Const("Trueprop",_) $ (Const("Not",_) $ Free(p,_))) (Const("Trueprop",_) $ Free(q,_)) = (p = q)
15371
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    63
  | is_neg _ _ = false;
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    64
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    65
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    66
exception STRIP_CONCL;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    67
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    68
15371
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    69
fun strip_concl' prems bvs (Const ("==>",_) $ P $ Q) =
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    70
      let val P' = HOLogic.dest_Trueprop P
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    71
  	  val prems' = P'::prems
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    72
      in
15371
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    73
	strip_concl' prems' bvs  Q
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    74
      end
15371
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    75
  | strip_concl' prems bvs P = 
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    76
      let val P' = HOLogic.Not $ (HOLogic.dest_Trueprop P)
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    77
      in
15371
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    78
	add_EX (make_conjs (P'::prems)) bvs
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    79
      end;
15371
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    80
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    81
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    82
fun strip_concl prems bvs concl (Const ("all", _) $ Abs (x,xtp,body))  = strip_concl prems ((x,xtp)::bvs) concl body
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    83
  | strip_concl prems bvs concl (Const ("==>",_) $ P $ Q) =
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    84
    if (is_neg P concl) then (strip_concl' prems bvs Q)
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    85
    else
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    86
	(let val P' = HOLogic.dest_Trueprop P
15371
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    87
	     val prems' = P'::prems
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    88
	 in
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    89
	     strip_concl prems' bvs  concl Q
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    90
	 end)
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    91
  | strip_concl prems bvs concl _ = add_EX (make_conjs prems) bvs;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    92
 
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    93
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    94
15371
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    95
fun trans_elim (main,others,concl) =
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    96
    let val others' = map (strip_concl [] [] concl) others
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    97
	val disjs = make_disjs others'
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    98
    in
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    99
	HOLogic.mk_imp (HOLogic.dest_Trueprop main, disjs)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   100
    end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   101
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   102
15390
87f78411f7c9 Comments and other tweaks by Jia
paulson
parents: 15371
diff changeset
   103
(* aux function of elim2Fol, take away predicate variable. *)
15371
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
   104
fun elimR2Fol_aux prems concl = 
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   105
    let val nprems = length prems
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   106
	val main = hd prems
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   107
    in
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   108
	if (nprems = 1) then HOLogic.Not $ (HOLogic.dest_Trueprop main)
15371
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
   109
        else trans_elim (main, tl prems, concl)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   110
    end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   111
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   112
    
16012
4ae42d8f2fea bug fixes for clause form transformation
paulson
parents: 16009
diff changeset
   113
(* convert an elim rule into an equivalent formula, of type term. *)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   114
fun elimR2Fol elimR = 
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   115
    let val elimR' = Drule.freeze_all elimR
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   116
	val (prems,concl) = (prems_of elimR', concl_of elimR')
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   117
    in
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   118
	case concl of Const("Trueprop",_) $ Free(_,Type("bool",[])) 
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   119
		      => HOLogic.mk_Trueprop (elimR2Fol_aux prems concl)
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   120
                    | Free(x,Type("prop",[])) => HOLogic.mk_Trueprop(elimR2Fol_aux prems concl) 
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   121
		    | _ => raise ELIMR2FOL("Not an elimination rule!")
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   122
    end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   123
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   124
15390
87f78411f7c9 Comments and other tweaks by Jia
paulson
parents: 15371
diff changeset
   125
(* check if a rule is an elim rule *)
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   126
fun is_elimR th = 
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   127
    case (concl_of th) of (Const ("Trueprop", _) $ Var (idx,_)) => true
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   128
			 | Var(indx,Type("prop",[])) => true
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   129
			 | _ => false;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   130
15997
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
   131
(* convert an elim-rule into an equivalent theorem that does not have the 
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
   132
   predicate variable.  Leave other theorems unchanged.*) 
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   133
fun transform_elim th =
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   134
  if is_elimR th then
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   135
    let val tm = elimR2Fol th
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   136
	val ctm = cterm_of (sign_of_thm th) tm	
15997
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
   137
    in
16039
dfe264950511 Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents: 16012
diff changeset
   138
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   139
	prove_goalw_cterm [] ctm (fn prems => [elimRule_tac th])
15997
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
   140
    end
16039
dfe264950511 Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents: 16012
diff changeset
   141
 else th;;	
dfe264950511 Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents: 16012
diff changeset
   142
dfe264950511 Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents: 16012
diff changeset
   143
dfe264950511 Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents: 16012
diff changeset
   144
15997
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
   145
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
   146
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
   147
(**** Transformation of Clasets and Simpsets into First-Order Axioms ****)
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
   148
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
   149
(* to be fixed: cnf_intro, cnf_rule, is_introR *)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   150
15390
87f78411f7c9 Comments and other tweaks by Jia
paulson
parents: 15371
diff changeset
   151
(* repeated resolution *)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   152
fun repeat_RS thm1 thm2 =
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   153
    let val thm1' =  thm1 RS thm2 handle THM _ => thm1
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   154
    in
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   155
	if eq_thm(thm1,thm1') then thm1' else (repeat_RS thm1' thm2)
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   156
    end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   157
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   158
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   159
(*Convert a theorem into NNF and also skolemize it. Original version, using
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   160
  Hilbert's epsilon in the resulting clauses.*)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   161
fun skolem_axiom th = 
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   162
  if Term.is_first_order (prop_of th) then
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   163
    let val th' = (skolemize o make_nnf o ObjectLogic.atomize_thm o Drule.freeze_all) th
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   164
    in 
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   165
	repeat_RS th' someI_ex
15872
8336ff711d80 fixed treatment of higher-order simprules
paulson
parents: 15736
diff changeset
   166
    end
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   167
  else raise THM ("skolem_axiom: not first-order", 0, [th]);
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   168
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   169
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   170
fun cnf_rule th = make_clauses [skolem_axiom (transform_elim th)];
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   171
15370
05b03ea0f18d trying to fix the transfer problem
paulson
parents: 15359
diff changeset
   172
(*Transfer a theorem in to theory Reconstruction.thy if it is not already
15359
8bad1f42fec0 new CLAUSIFY attribute for proof reconstruction with lemmas
paulson
parents: 15347
diff changeset
   173
  inside that theory -- because it's needed for Skolemization *)
8bad1f42fec0 new CLAUSIFY attribute for proof reconstruction with lemmas
paulson
parents: 15347
diff changeset
   174
15370
05b03ea0f18d trying to fix the transfer problem
paulson
parents: 15359
diff changeset
   175
val recon_thy = ThyInfo.get_theory"Reconstruction";
15359
8bad1f42fec0 new CLAUSIFY attribute for proof reconstruction with lemmas
paulson
parents: 15347
diff changeset
   176
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   177
fun transfer_to_Reconstruction th =
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   178
    transfer recon_thy th handle THM _ => th;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   179
15955
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   180
fun is_taut th =
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   181
      case (prop_of th) of
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   182
           (Const ("Trueprop", _) $ Const ("True", _)) => true
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   183
         | _ => false;
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   184
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   185
(* remove tautologous clauses *)
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   186
val rm_redundant_cls = List.filter (not o is_taut);
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   187
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   188
(* transform an Isabelle thm into CNF *)
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   189
fun cnf_axiom_aux th =
16173
9e2f6c0a779d clausification bug fix
paulson
parents: 16156
diff changeset
   190
    map zero_var_indexes
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   191
        (rm_redundant_cls (cnf_rule (transfer_to_Reconstruction th)));
15997
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
   192
       
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
   193
       
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   194
(**** SKOLEMIZATION BY INFERENCE (lcp) ****)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   195
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   196
(*Traverse a term, accumulating Skolem function definitions.*)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   197
fun declare_skofuns s t thy =
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   198
  let fun dec_sko (Const ("Ex",_) $ (xtp as Abs(_,T,p))) (n, thy) =
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   199
	    (*Existential: declare a Skolem function, then insert into body and continue*)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   200
	    let val cname = s ^ "_" ^ Int.toString n
16012
4ae42d8f2fea bug fixes for clause form transformation
paulson
parents: 16009
diff changeset
   201
		val args = term_frees xtp  (*get the formal parameter list*)
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   202
		val Ts = map type_of args
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   203
		val cT = Ts ---> T
16125
bd13a0017137 proper use of Sign.full_name;
wenzelm
parents: 16039
diff changeset
   204
		val c = Const (Sign.full_name (Theory.sign_of thy) cname, cT)
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   205
		val rhs = list_abs_free (map dest_Free args, HOLogic.choice_const T $ xtp)
16012
4ae42d8f2fea bug fixes for clause form transformation
paulson
parents: 16009
diff changeset
   206
		        (*Forms a lambda-abstraction over the formal parameters*)
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   207
		val def = equals cT $ c $ rhs
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   208
		val thy' = Theory.add_consts_i [(cname, cT, NoSyn)] thy
16012
4ae42d8f2fea bug fixes for clause form transformation
paulson
parents: 16009
diff changeset
   209
		           (*Theory is augmented with the constant, then its def*)
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   210
		val thy'' = Theory.add_defs_i false [(cname ^ "_def", def)] thy'
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   211
	    in dec_sko (subst_bound (list_comb(c,args), p)) (n+1, thy'') end
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   212
	| dec_sko (Const ("All",_) $ (xtp as Abs(a,T,p))) (n, thy) =
16012
4ae42d8f2fea bug fixes for clause form transformation
paulson
parents: 16009
diff changeset
   213
	    (*Universal quant: insert a free variable into body and continue*)
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   214
	    let val fname = variant (add_term_names (p,[])) a
16012
4ae42d8f2fea bug fixes for clause form transformation
paulson
parents: 16009
diff changeset
   215
	    in dec_sko (subst_bound (Free(fname,T), p)) (n, thy) end
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   216
	| dec_sko (Const ("op &", _) $ p $ q) nthy = 
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   217
	    dec_sko q (dec_sko p nthy)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   218
	| dec_sko (Const ("op |", _) $ p $ q) nthy = 
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   219
	    dec_sko q (dec_sko p nthy)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   220
	| dec_sko (Const ("Trueprop", _) $ p) nthy = 
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   221
	    dec_sko p nthy
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   222
	| dec_sko t (n,thy) = (n,thy) (*Do nothing otherwise*)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   223
  in  #2 (dec_sko t (1,thy))  end;
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   224
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   225
(*cterms are used throughout for efficiency*)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   226
val cTrueprop = Thm.cterm_of (Theory.sign_of HOL.thy) HOLogic.Trueprop;
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   227
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   228
(*cterm version of mk_cTrueprop*)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   229
fun c_mkTrueprop A = Thm.capply cTrueprop A;
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   230
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   231
(*Given an abstraction over n variables, replace the bound variables by free
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   232
  ones. Return the body, along with the list of free variables.*)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   233
fun c_variant_abs_multi (ct0, vars) = 
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   234
      let val (cv,ct) = Thm.dest_abs NONE ct0
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   235
      in  c_variant_abs_multi (ct, cv::vars)  end
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   236
      handle CTERM _ => (ct0, rev vars);
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   237
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   238
(*Given the definition of a Skolem function, return a theorem to replace 
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   239
  an existential formula by a use of that function.*)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   240
fun skolem_of_def def =
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   241
  let val (c,rhs) = Drule.dest_equals (cprop_of (Drule.freeze_all def))
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   242
      val (ch, frees) = c_variant_abs_multi (rhs, [])
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   243
      val (chil,cabs) = Thm.dest_comb ch
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   244
      val {sign, t, ...} = rep_cterm chil
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   245
      val (Const ("Hilbert_Choice.Eps", Type("fun",[_,T]))) = t
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   246
      val cex = Thm.cterm_of sign (HOLogic.exists_const T)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   247
      val ex_tm = c_mkTrueprop (Thm.capply cex cabs)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   248
      and conc =  c_mkTrueprop (Drule.beta_conv cabs (Drule.list_comb(c,frees)));
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   249
  in  prove_goalw_cterm [def] (Drule.mk_implies (ex_tm, conc))
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   250
	    (fn [prem] => [ rtac (prem RS someI_ex) 1 ])
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   251
  end;	 
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   252
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   253
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   254
(*Converts an Isabelle theorem (intro, elim or simp format) into nnf.*)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   255
fun to_nnf thy th = 
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   256
    if Term.is_first_order (prop_of th) then
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   257
      th |> Thm.transfer thy |> transform_elim |> Drule.freeze_all
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   258
	 |> ObjectLogic.atomize_thm |> make_nnf
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   259
    else raise THM ("to_nnf: not first-order", 0, [th]);
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   260
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   261
(*The cache prevents repeated clausification of a theorem, 
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   262
  and also repeated declaration of Skolem functions*)  
15955
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   263
val clause_cache = ref (Symtab.empty : (thm * thm list) Symtab.table)
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   264
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   265
(*Declare Skolem functions for a theorem, supplied in nnf and with its name*)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   266
fun skolem thy (name,th) =
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   267
  let val cname = (case name of
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   268
                         "" => gensym "sko" | s => Sign.base_name s)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   269
      val thy' = declare_skofuns cname (#prop (rep_thm th)) thy
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   270
  in (map (skolem_of_def o #2) (axioms_of thy'), thy') end;
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   271
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   272
(*Populate the clause cache using the supplied theorems*)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   273
fun skolemlist [] thy = thy
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   274
  | skolemlist ((name,th)::nths) thy = 
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   275
      (case Symtab.lookup (!clause_cache,name) of
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   276
	  NONE => 
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   277
	    let val nnfth = to_nnf thy th
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   278
		val (skoths,thy') = skolem thy (name, nnfth)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   279
		val cls = Meson.make_cnf skoths nnfth
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   280
	    in  clause_cache := Symtab.update ((name, (th,cls)), !clause_cache);
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   281
		skolemlist nths thy'
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   282
	    end
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   283
	| SOME _ => skolemlist nths thy) (*FIXME: check for duplicate names?*)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   284
      handle THM _ => skolemlist nths thy;
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   285
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   286
(*Exported function to convert Isabelle theorems into axiom clauses*) 
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   287
fun cnf_axiom (name,th) =
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   288
    case name of
15955
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   289
	  "" => cnf_axiom_aux th (*no name, so can't cache*)
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   290
	| s  => case Symtab.lookup (!clause_cache,s) of
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   291
	  	  NONE => 
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   292
		    let val cls = cnf_axiom_aux th
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   293
		    in  clause_cache := Symtab.update ((s, (th,cls)), !clause_cache); cls
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   294
		    end
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   295
	        | SOME(th',cls) =>
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   296
		    if eq_thm(th,th') then cls
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   297
		    else (*New theorem stored under the same name? Possible??*)
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   298
		      let val cls = cnf_axiom_aux th
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   299
		      in  clause_cache := Symtab.update ((s, (th,cls)), !clause_cache); cls
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   300
		      end;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   301
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   302
fun pairname th = (Thm.name_of_thm th, th);
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   303
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   304
fun meta_cnf_axiom th = 
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   305
    map Meson.make_meta_clause (cnf_axiom (pairname th));
15499
419dc5ffe8bc clausification and proof reconstruction
paulson
parents: 15495
diff changeset
   306
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   307
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   308
(* changed: with one extra case added *)
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   309
fun univ_vars_of_aux (Const ("Hilbert_Choice.Eps",_) $ Abs(_,_,body)) vars =    
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   310
      univ_vars_of_aux body vars
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   311
  | univ_vars_of_aux (Const ("Ex",_) $ Abs(_,_,body)) vars = 
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   312
      univ_vars_of_aux body vars (* EX x. body *)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   313
  | univ_vars_of_aux (P $ Q) vars =
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   314
      univ_vars_of_aux Q (univ_vars_of_aux P vars)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   315
  | univ_vars_of_aux (t as Var(_,_)) vars = 
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   316
      if (t mem vars) then vars else (t::vars)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   317
  | univ_vars_of_aux _ vars = vars;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   318
  
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   319
fun univ_vars_of t = univ_vars_of_aux t [];
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   320
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   321
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   322
fun get_new_skolem epss (t as (Const ("Hilbert_Choice.Eps",_) $ Abs(_,tp,_)))  = 
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   323
    let val all_vars = univ_vars_of t
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   324
	val sk_term = ResSkolemFunction.gen_skolem all_vars tp
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   325
    in
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   326
	(sk_term,(t,sk_term)::epss)
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   327
    end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   328
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   329
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15499
diff changeset
   330
fun sk_lookup [] t = NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15499
diff changeset
   331
  | sk_lookup ((tm,sk_tm)::tms) t = if (t = tm) then SOME (sk_tm) else (sk_lookup tms t);
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   332
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   333
15390
87f78411f7c9 Comments and other tweaks by Jia
paulson
parents: 15371
diff changeset
   334
87f78411f7c9 Comments and other tweaks by Jia
paulson
parents: 15371
diff changeset
   335
(* get the proper skolem term to replace epsilon term *)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   336
fun get_skolem epss t = 
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   337
    case (sk_lookup epss t) of NONE => get_new_skolem epss t
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   338
		             | SOME sk => (sk,epss);
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   339
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   340
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   341
fun rm_Eps_cls_aux epss (t as (Const ("Hilbert_Choice.Eps",_) $ Abs(_,_,_))) = 
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   342
       get_skolem epss t
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   343
  | rm_Eps_cls_aux epss (P $ Q) =
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   344
       let val (P',epss') = rm_Eps_cls_aux epss P
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   345
	   val (Q',epss'') = rm_Eps_cls_aux epss' Q
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   346
       in (P' $ Q',epss'') end
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   347
  | rm_Eps_cls_aux epss t = (t,epss);
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   348
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   349
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   350
fun rm_Eps_cls epss th = rm_Eps_cls_aux epss (prop_of th);
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   351
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   352
15390
87f78411f7c9 Comments and other tweaks by Jia
paulson
parents: 15371
diff changeset
   353
(* remove the epsilon terms in a formula, by skolem terms. *)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   354
fun rm_Eps _ [] = []
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   355
  | rm_Eps epss (th::thms) = 
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   356
      let val (th',epss') = rm_Eps_cls epss th
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   357
      in th' :: (rm_Eps epss' thms) end;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   358
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   359
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   360
15872
8336ff711d80 fixed treatment of higher-order simprules
paulson
parents: 15736
diff changeset
   361
(**** Extract and Clausify theorems from a theory's claset and simpset ****)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   362
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   363
fun claset_rules_of_thy thy =
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   364
    let val clsset = rep_cs (claset_of thy)
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   365
	val safeEs = #safeEs clsset
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   366
	val safeIs = #safeIs clsset
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   367
	val hazEs = #hazEs clsset
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   368
	val hazIs = #hazIs clsset
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   369
    in
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   370
	map pairname (safeEs @ safeIs @ hazEs @ hazIs)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   371
    end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   372
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   373
fun simpset_rules_of_thy thy =
15872
8336ff711d80 fixed treatment of higher-order simprules
paulson
parents: 15736
diff changeset
   374
    let val rules = #rules(fst (rep_ss (simpset_of thy)))
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   375
    in
15872
8336ff711d80 fixed treatment of higher-order simprules
paulson
parents: 15736
diff changeset
   376
	map (fn (_,r) => (#name r, #thm r)) (Net.dest rules)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   377
    end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   378
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   379
15872
8336ff711d80 fixed treatment of higher-order simprules
paulson
parents: 15736
diff changeset
   380
(**** Translate a set of classical/simplifier rules into CNF (still as type "thm")  ****)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   381
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   382
(* classical rules *)
15872
8336ff711d80 fixed treatment of higher-order simprules
paulson
parents: 15736
diff changeset
   383
fun cnf_rules [] err_list = ([],err_list)
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   384
  | cnf_rules ((name,th) :: thms) err_list = 
15872
8336ff711d80 fixed treatment of higher-order simprules
paulson
parents: 15736
diff changeset
   385
      let val (ts,es) = cnf_rules thms err_list
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   386
      in  (cnf_axiom (name,th) :: ts,es) handle  _ => (ts, (th::es))  end;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   387
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   388
(* CNF all rules from a given theory's classical reasoner *)
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   389
fun cnf_classical_rules_thy thy = 
15872
8336ff711d80 fixed treatment of higher-order simprules
paulson
parents: 15736
diff changeset
   390
    cnf_rules (claset_rules_of_thy thy) [];
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   391
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   392
(* CNF all simplifier rules from a given theory's simpset *)
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   393
fun cnf_simpset_rules_thy thy =
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   394
    cnf_rules (simpset_rules_of_thy thy) [];
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   395
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   396
15872
8336ff711d80 fixed treatment of higher-order simprules
paulson
parents: 15736
diff changeset
   397
(**** Convert all theorems of a claset/simpset into clauses (ResClause.clause) ****)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   398
16156
2f6fc19aba1e Fixed array containing clasimpset rules. Added flags to turn on and off reconstruction and full spass
quigley
parents: 16125
diff changeset
   399
(* outputs a list of (clause,thm) pairs *)
2f6fc19aba1e Fixed array containing clasimpset rules. Added flags to turn on and off reconstruction and full spass
quigley
parents: 16125
diff changeset
   400
(*fun clausify_axiom_pairs (thm_name,thm) =
2f6fc19aba1e Fixed array containing clasimpset rules. Added flags to turn on and off reconstruction and full spass
quigley
parents: 16125
diff changeset
   401
    let val isa_clauses = cnf_axiom (thm_name,thm) (*"isa_clauses" are already "standard"ed. *)
2f6fc19aba1e Fixed array containing clasimpset rules. Added flags to turn on and off reconstruction and full spass
quigley
parents: 16125
diff changeset
   402
        val isa_clauses' = rm_Eps [] isa_clauses
2f6fc19aba1e Fixed array containing clasimpset rules. Added flags to turn on and off reconstruction and full spass
quigley
parents: 16125
diff changeset
   403
        val clauses_n = length isa_clauses
2f6fc19aba1e Fixed array containing clasimpset rules. Added flags to turn on and off reconstruction and full spass
quigley
parents: 16125
diff changeset
   404
	fun make_axiom_clauses _ [] []= []
2f6fc19aba1e Fixed array containing clasimpset rules. Added flags to turn on and off reconstruction and full spass
quigley
parents: 16125
diff changeset
   405
	  | make_axiom_clauses i (cls::clss) (cls'::clss')= ((ResClause.make_axiom_clause cls (thm_name,i)),cls', thm_name, i) :: make_axiom_clauses (i+1) clss clss'
2f6fc19aba1e Fixed array containing clasimpset rules. Added flags to turn on and off reconstruction and full spass
quigley
parents: 16125
diff changeset
   406
    in
2f6fc19aba1e Fixed array containing clasimpset rules. Added flags to turn on and off reconstruction and full spass
quigley
parents: 16125
diff changeset
   407
	make_axiom_clauses 0 isa_clauses' isa_clauses		
2f6fc19aba1e Fixed array containing clasimpset rules. Added flags to turn on and off reconstruction and full spass
quigley
parents: 16125
diff changeset
   408
    end;
2f6fc19aba1e Fixed array containing clasimpset rules. Added flags to turn on and off reconstruction and full spass
quigley
parents: 16125
diff changeset
   409
*)
2f6fc19aba1e Fixed array containing clasimpset rules. Added flags to turn on and off reconstruction and full spass
quigley
parents: 16125
diff changeset
   410
2f6fc19aba1e Fixed array containing clasimpset rules. Added flags to turn on and off reconstruction and full spass
quigley
parents: 16125
diff changeset
   411
16039
dfe264950511 Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents: 16012
diff changeset
   412
fun clausify_axiom_pairs (thm_name,thm) =
dfe264950511 Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents: 16012
diff changeset
   413
    let val isa_clauses = cnf_axiom (thm_name,thm) (*"isa_clauses" are already "standard"ed. *)
dfe264950511 Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents: 16012
diff changeset
   414
        val isa_clauses' = rm_Eps [] isa_clauses
dfe264950511 Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents: 16012
diff changeset
   415
        val clauses_n = length isa_clauses
dfe264950511 Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents: 16012
diff changeset
   416
	fun make_axiom_clauses _ [] []= []
dfe264950511 Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents: 16012
diff changeset
   417
	  | make_axiom_clauses i (cls::clss) (cls'::clss')= ((ResClause.make_axiom_clause cls (thm_name,i)),cls') :: make_axiom_clauses (i+1) clss clss'
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   418
    in
16039
dfe264950511 Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents: 16012
diff changeset
   419
	make_axiom_clauses 0 isa_clauses' isa_clauses		
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   420
    end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   421
16039
dfe264950511 Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents: 16012
diff changeset
   422
fun clausify_rules_pairs [] err_list = ([],err_list)
dfe264950511 Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents: 16012
diff changeset
   423
  | clausify_rules_pairs ((name,thm)::thms) err_list =
dfe264950511 Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents: 16012
diff changeset
   424
    let val (ts,es) = clausify_rules_pairs thms err_list
dfe264950511 Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents: 16012
diff changeset
   425
    in
dfe264950511 Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents: 16012
diff changeset
   426
	((clausify_axiom_pairs (name,thm))::ts,es) handle  _ => (ts,(thm::es))
dfe264950511 Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents: 16012
diff changeset
   427
    end;
dfe264950511 Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents: 16012
diff changeset
   428
(* classical rules *)
dfe264950511 Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents: 16012
diff changeset
   429
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   430
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   431
(*Setup function: takes a theory and installs ALL simprules and claset rules 
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   432
  into the clause cache*)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   433
fun clause_cache_setup thy =
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   434
  let val simps = simpset_rules_of_thy thy
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   435
      and clas  = claset_rules_of_thy thy
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   436
  in skolemlist clas (skolemlist simps thy) end;
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   437
  
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   438
val setup = [clause_cache_setup];  
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   439
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   440
end;