src/HOL/Tools/res_axioms.ML
author mengj
Fri, 28 Oct 2005 02:25:57 +0200
changeset 18000 ac059afd6b86
parent 17959 8db36a108213
child 18009 df077e122064
permissions -rw-r--r--
Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names. Also added function "repeat_RS" to the signature.
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
17404
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
    11
  val tagging_enabled : bool
15997
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
    12
  val elimRule_tac : thm -> Tactical.tactic
16012
4ae42d8f2fea bug fixes for clause form transformation
paulson
parents: 16009
diff changeset
    13
  val elimR2Fol : thm -> term
15997
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
    14
  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
    15
  val clausify_axiom_pairs : (string*thm) -> (ResClause.clause*thm) list
18000
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
    16
  val clausify_axiom_pairsH : (string*thm) -> (ResHolClause.clause*thm) list
15997
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
    17
  val cnf_axiom : (string * thm) -> thm list
18000
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
    18
  val cnf_axiomH : (string * thm) -> thm list
15997
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
    19
  val meta_cnf_axiom : thm -> thm list
18000
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
    20
  val meta_cnf_axiomH : thm -> thm list
16012
4ae42d8f2fea bug fixes for clause form transformation
paulson
parents: 16009
diff changeset
    21
  val rm_Eps : (term * term) list -> thm list -> term list
15997
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
    22
  val claset_rules_of_thy : theory -> (string * thm) list
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
    23
  val simpset_rules_of_thy : theory -> (string * thm) list
17484
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17412
diff changeset
    24
  val claset_rules_of_ctxt: Proof.context -> (string * thm) list
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17412
diff changeset
    25
  val simpset_rules_of_ctxt : Proof.context -> (string * thm) list
17829
35123f89801e simplifying the treatment of clausification
paulson
parents: 17819
diff changeset
    26
  val clausify_rules_pairs : (string*thm) list -> (ResClause.clause*thm) list list
18000
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
    27
  val clausify_rules_pairsH : (string*thm) list -> (ResHolClause.clause*thm) list list
16563
a92f96951355 meson method taking an argument list
paulson
parents: 16173
diff changeset
    28
  val clause_setup : (theory -> theory) list
a92f96951355 meson method taking an argument list
paulson
parents: 16173
diff changeset
    29
  val meson_method_setup : (theory -> theory) list
17905
1574533861b1 Added files in order to use external ATPs as oracles and invoke these ATPs by calling Isabelle methods (currently "vampire" and "eprover").
mengj
parents: 17829
diff changeset
    30
  val pairname : thm -> (string * thm)
18000
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
    31
  val repeat_RS : thm -> thm -> thm
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
    32
15997
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
    33
  end;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    34
15997
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
    35
structure ResAxioms : RES_AXIOMS =
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
    36
 
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
    37
struct
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    38
18000
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
    39
17404
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
    40
val tagging_enabled = false (*compile_time option*)
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
    41
15997
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
    42
(**** Transformation of Elimination Rules into First-Order Formulas****)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    43
15390
87f78411f7c9 Comments and other tweaks by Jia
paulson
parents: 15371
diff changeset
    44
(* a tactic used to prove an elim-rule. *)
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
    45
fun elimRule_tac th =
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
    46
    ((rtac impI 1) ORELSE (rtac notI 1)) THEN (etac th 1) THEN
16588
8de758143786 stricter first-order check for meson
paulson
parents: 16563
diff changeset
    47
    REPEAT(fast_tac HOL_cs 1);
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    48
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    49
exception ELIMR2FOL of string;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    50
15390
87f78411f7c9 Comments and other tweaks by Jia
paulson
parents: 15371
diff changeset
    51
(* functions used to construct a formula *)
87f78411f7c9 Comments and other tweaks by Jia
paulson
parents: 15371
diff changeset
    52
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    53
fun make_disjs [x] = x
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    54
  | make_disjs (x :: xs) = HOLogic.mk_disj(x, make_disjs xs)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    55
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    56
fun make_conjs [x] = x
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    57
  | make_conjs (x :: xs) =  HOLogic.mk_conj(x, make_conjs xs)
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    58
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    59
fun add_EX tm [] = tm
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    60
  | 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
    61
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    62
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    63
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    64
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
    65
  | is_neg _ _ = false;
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    66
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    67
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    68
exception STRIP_CONCL;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    69
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    70
15371
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    71
fun strip_concl' prems bvs (Const ("==>",_) $ P $ Q) =
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    72
      let val P' = HOLogic.dest_Trueprop P
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    73
  	  val prems' = P'::prems
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    74
      in
15371
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    75
	strip_concl' prems' bvs  Q
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    76
      end
15371
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    77
  | strip_concl' prems bvs P = 
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    78
      let val P' = HOLogic.Not $ (HOLogic.dest_Trueprop P)
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    79
      in
15371
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    80
	add_EX (make_conjs (P'::prems)) bvs
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    81
      end;
15371
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    82
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    83
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    84
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
    85
  | strip_concl prems bvs concl (Const ("==>",_) $ P $ Q) =
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    86
    if (is_neg P concl) then (strip_concl' prems bvs Q)
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    87
    else
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    88
	(let val P' = HOLogic.dest_Trueprop P
15371
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    89
	     val prems' = P'::prems
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    90
	 in
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    91
	     strip_concl prems' bvs  concl Q
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    92
	 end)
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    93
  | strip_concl prems bvs concl _ = add_EX (make_conjs prems) bvs;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    94
 
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    95
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    96
15371
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    97
fun trans_elim (main,others,concl) =
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    98
    let val others' = map (strip_concl [] [] concl) others
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    99
	val disjs = make_disjs others'
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   100
    in
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   101
	HOLogic.mk_imp (HOLogic.dest_Trueprop main, disjs)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   102
    end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   103
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   104
15390
87f78411f7c9 Comments and other tweaks by Jia
paulson
parents: 15371
diff changeset
   105
(* aux function of elim2Fol, take away predicate variable. *)
15371
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
   106
fun elimR2Fol_aux prems concl = 
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   107
    let val nprems = length prems
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   108
	val main = hd prems
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   109
    in
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   110
	if (nprems = 1) then HOLogic.Not $ (HOLogic.dest_Trueprop main)
15371
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
   111
        else trans_elim (main, tl prems, concl)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   112
    end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   113
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   114
    
16012
4ae42d8f2fea bug fixes for clause form transformation
paulson
parents: 16009
diff changeset
   115
(* convert an elim rule into an equivalent formula, of type term. *)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   116
fun elimR2Fol elimR = 
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   117
    let val elimR' = Drule.freeze_all elimR
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   118
	val (prems,concl) = (prems_of elimR', concl_of elimR')
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   119
    in
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   120
	case concl of Const("Trueprop",_) $ Free(_,Type("bool",[])) 
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   121
		      => HOLogic.mk_Trueprop (elimR2Fol_aux prems concl)
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   122
                    | Free(x,Type("prop",[])) => HOLogic.mk_Trueprop(elimR2Fol_aux prems concl) 
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   123
		    | _ => raise ELIMR2FOL("Not an elimination rule!")
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   124
    end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   125
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   126
15390
87f78411f7c9 Comments and other tweaks by Jia
paulson
parents: 15371
diff changeset
   127
(* check if a rule is an elim rule *)
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   128
fun is_elimR th = 
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   129
    case (concl_of th) of (Const ("Trueprop", _) $ Var (idx,_)) => true
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   130
			 | Var(indx,Type("prop",[])) => true
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   131
			 | _ => false;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   132
15997
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
   133
(* convert an elim-rule into an equivalent theorem that does not have the 
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
   134
   predicate variable.  Leave other theorems unchanged.*) 
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   135
fun transform_elim th =
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   136
  if is_elimR th then
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   137
    let val tm = elimR2Fol th
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   138
	val ctm = cterm_of (sign_of_thm th) tm	
15997
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
   139
    in
17959
8db36a108213 OldGoals;
wenzelm
parents: 17905
diff changeset
   140
	OldGoals.prove_goalw_cterm [] ctm (fn prems => [elimRule_tac th])
15997
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
   141
    end
16563
a92f96951355 meson method taking an argument list
paulson
parents: 16173
diff changeset
   142
 else th;
15997
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
   143
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
   144
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
   145
(**** Transformation of Clasets and Simpsets into First-Order Axioms ****)
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
   146
15390
87f78411f7c9 Comments and other tweaks by Jia
paulson
parents: 15371
diff changeset
   147
(* repeated resolution *)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   148
fun repeat_RS thm1 thm2 =
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   149
    let val thm1' =  thm1 RS thm2 handle THM _ => thm1
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   150
    in
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   151
	if eq_thm(thm1,thm1') then thm1' else (repeat_RS thm1' thm2)
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   152
    end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   153
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   154
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   155
(*Convert a theorem into NNF and also skolemize it. Original version, using
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   156
  Hilbert's epsilon in the resulting clauses.*)
18000
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   157
fun skolem_axiom_g mk_nnf th = 
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   158
  let val th' = (skolemize o mk_nnf o ObjectLogic.atomize_thm o Drule.freeze_all) th
16588
8de758143786 stricter first-order check for meson
paulson
parents: 16563
diff changeset
   159
  in  repeat_RS th' someI_ex
8de758143786 stricter first-order check for meson
paulson
parents: 16563
diff changeset
   160
  end;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   161
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   162
18000
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   163
val skolem_axiom = skolem_axiom_g make_nnf;
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   164
val skolem_axiomH = skolem_axiom_g make_nnf1;
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   165
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   166
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   167
fun cnf_rule th = make_clauses [skolem_axiom (transform_elim th)];
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   168
18000
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   169
fun cnf_ruleH th = make_clauses [skolem_axiomH (transform_elim th)];
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   170
16563
a92f96951355 meson method taking an argument list
paulson
parents: 16173
diff changeset
   171
(*Transfer a theorem into theory Reconstruction.thy if it is not already
15359
8bad1f42fec0 new CLAUSIFY attribute for proof reconstruction with lemmas
paulson
parents: 15347
diff changeset
   172
  inside that theory -- because it's needed for Skolemization *)
8bad1f42fec0 new CLAUSIFY attribute for proof reconstruction with lemmas
paulson
parents: 15347
diff changeset
   173
16563
a92f96951355 meson method taking an argument list
paulson
parents: 16173
diff changeset
   174
(*This will refer to the final version of theory Reconstruction.*)
a92f96951355 meson method taking an argument list
paulson
parents: 16173
diff changeset
   175
val recon_thy_ref = Theory.self_ref (the_context ());  
15359
8bad1f42fec0 new CLAUSIFY attribute for proof reconstruction with lemmas
paulson
parents: 15347
diff changeset
   176
16563
a92f96951355 meson method taking an argument list
paulson
parents: 16173
diff changeset
   177
(*If called while Reconstruction is being created, it will transfer to the
a92f96951355 meson method taking an argument list
paulson
parents: 16173
diff changeset
   178
  current version. If called afterward, it will transfer to the final version.*)
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   179
fun transfer_to_Reconstruction th =
16563
a92f96951355 meson method taking an argument list
paulson
parents: 16173
diff changeset
   180
    transfer (Theory.deref recon_thy_ref) th handle THM _ => th;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   181
15955
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   182
fun is_taut th =
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   183
      case (prop_of th) of
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   184
           (Const ("Trueprop", _) $ Const ("True", _)) => true
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   185
         | _ => false;
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   186
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   187
(* remove tautologous clauses *)
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   188
val rm_redundant_cls = List.filter (not o is_taut);
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   189
17829
35123f89801e simplifying the treatment of clausification
paulson
parents: 17819
diff changeset
   190
(* transform an Isabelle theorem into CNF *)
18000
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   191
fun cnf_axiom_aux_g cnf_rule th =
16173
9e2f6c0a779d clausification bug fix
paulson
parents: 16156
diff changeset
   192
    map zero_var_indexes
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   193
        (rm_redundant_cls (cnf_rule (transfer_to_Reconstruction th)));
18000
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   194
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   195
val cnf_axiom_aux = cnf_axiom_aux_g cnf_rule;
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   196
val cnf_axiom_auxH = cnf_axiom_aux_g cnf_ruleH;
15997
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
   197
       
c71031d7988c consolidation and simplification
paulson
parents: 15956
diff changeset
   198
       
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   199
(**** SKOLEMIZATION BY INFERENCE (lcp) ****)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   200
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   201
(*Traverse a term, accumulating Skolem function definitions.*)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   202
fun declare_skofuns s t thy =
17404
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   203
  let fun dec_sko (Const ("Ex",_) $ (xtp as Abs(_,T,p))) (n, (thy, axs)) =
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   204
	    (*Existential: declare a Skolem function, then insert into body and continue*)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   205
	    let val cname = s ^ "_" ^ Int.toString n
16012
4ae42d8f2fea bug fixes for clause form transformation
paulson
parents: 16009
diff changeset
   206
		val args = term_frees xtp  (*get the formal parameter list*)
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   207
		val Ts = map type_of args
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   208
		val cT = Ts ---> T
16125
bd13a0017137 proper use of Sign.full_name;
wenzelm
parents: 16039
diff changeset
   209
		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
   210
		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
   211
		        (*Forms a lambda-abstraction over the formal parameters*)
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   212
		val def = equals cT $ c $ rhs
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   213
		val thy' = Theory.add_consts_i [(cname, cT, NoSyn)] thy
16012
4ae42d8f2fea bug fixes for clause form transformation
paulson
parents: 16009
diff changeset
   214
		           (*Theory is augmented with the constant, then its def*)
17404
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   215
		val cdef = cname ^ "_def"
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   216
		val thy'' = Theory.add_defs_i false [(cdef, def)] thy'
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   217
	    in dec_sko (subst_bound (list_comb(c,args), p)) 
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   218
	               (n+1, (thy'', get_axiom thy'' cdef :: axs)) 
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   219
	    end
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   220
	| dec_sko (Const ("All",_) $ (xtp as Abs(a,T,p))) (n, thx) =
16012
4ae42d8f2fea bug fixes for clause form transformation
paulson
parents: 16009
diff changeset
   221
	    (*Universal quant: insert a free variable into body and continue*)
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   222
	    let val fname = variant (add_term_names (p,[])) a
17404
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   223
	    in dec_sko (subst_bound (Free(fname,T), p)) (n, thx) end
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   224
	| dec_sko (Const ("op &", _) $ p $ q) nthy = 
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   225
	    dec_sko q (dec_sko p nthy)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   226
	| dec_sko (Const ("op |", _) $ p $ q) nthy = 
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   227
	    dec_sko q (dec_sko p nthy)
17404
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   228
	| dec_sko (Const ("HOL.tag", _) $ p) nthy = 
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   229
	    dec_sko p nthy
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   230
	| dec_sko (Const ("Trueprop", _) $ p) nthy = 
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   231
	    dec_sko p nthy
17404
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   232
	| dec_sko t nthx = nthx (*Do nothing otherwise*)
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   233
  in  #2 (dec_sko t (1, (thy,[])))  end;
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   234
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   235
(*cterms are used throughout for efficiency*)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   236
val cTrueprop = Thm.cterm_of (Theory.sign_of HOL.thy) HOLogic.Trueprop;
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
(*cterm version of mk_cTrueprop*)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   239
fun c_mkTrueprop A = Thm.capply cTrueprop A;
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   240
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   241
(*Given an abstraction over n variables, replace the bound variables by free
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   242
  ones. Return the body, along with the list of free variables.*)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   243
fun c_variant_abs_multi (ct0, vars) = 
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   244
      let val (cv,ct) = Thm.dest_abs NONE ct0
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   245
      in  c_variant_abs_multi (ct, cv::vars)  end
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   246
      handle CTERM _ => (ct0, rev vars);
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   247
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   248
(*Given the definition of a Skolem function, return a theorem to replace 
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   249
  an existential formula by a use of that function.*)
16588
8de758143786 stricter first-order check for meson
paulson
parents: 16563
diff changeset
   250
fun skolem_of_def def =  
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   251
  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
   252
      val (ch, frees) = c_variant_abs_multi (rhs, [])
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   253
      val (chil,cabs) = Thm.dest_comb ch
16588
8de758143786 stricter first-order check for meson
paulson
parents: 16563
diff changeset
   254
      val {sign,t, ...} = rep_cterm chil
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   255
      val (Const ("Hilbert_Choice.Eps", Type("fun",[_,T]))) = t
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   256
      val cex = Thm.cterm_of sign (HOLogic.exists_const T)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   257
      val ex_tm = c_mkTrueprop (Thm.capply cex cabs)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   258
      and conc =  c_mkTrueprop (Drule.beta_conv cabs (Drule.list_comb(c,frees)));
17959
8db36a108213 OldGoals;
wenzelm
parents: 17905
diff changeset
   259
  in  OldGoals.prove_goalw_cterm [def] (Drule.mk_implies (ex_tm, conc))
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   260
	    (fn [prem] => [ rtac (prem RS someI_ex) 1 ])
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   261
  end;	 
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   262
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   263
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   264
(*Converts an Isabelle theorem (intro, elim or simp format) into nnf.*)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   265
fun to_nnf thy th = 
16588
8de758143786 stricter first-order check for meson
paulson
parents: 16563
diff changeset
   266
    th |> Thm.transfer thy
8de758143786 stricter first-order check for meson
paulson
parents: 16563
diff changeset
   267
       |> transform_elim |> Drule.freeze_all
8de758143786 stricter first-order check for meson
paulson
parents: 16563
diff changeset
   268
       |> ObjectLogic.atomize_thm |> make_nnf;
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   269
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   270
(*The cache prevents repeated clausification of a theorem, 
16800
90eff1b52428 improved Net interface;
wenzelm
parents: 16588
diff changeset
   271
  and also repeated declaration of Skolem functions*)  (* FIXME better use Termtab!? *)
15955
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   272
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
   273
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   274
(*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
   275
fun skolem thy (name,th) =
16588
8de758143786 stricter first-order check for meson
paulson
parents: 16563
diff changeset
   276
  let val cname = (case name of "" => gensym "sko" | s => Sign.base_name s)
17404
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   277
      val (thy',axs) = declare_skofuns cname (#prop (rep_thm th)) thy
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   278
  in (map skolem_of_def axs, thy') end;
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   279
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   280
(*Populate the clause cache using the supplied theorems*)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   281
fun skolemlist [] thy = thy
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   282
  | skolemlist ((name,th)::nths) thy = 
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17404
diff changeset
   283
      (case Symtab.lookup (!clause_cache) name of
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   284
	  NONE => 
16588
8de758143786 stricter first-order check for meson
paulson
parents: 16563
diff changeset
   285
	    let val (nnfth,ok) = (to_nnf thy th, true)  
8de758143786 stricter first-order check for meson
paulson
parents: 16563
diff changeset
   286
	                 handle THM _ => (asm_rl, false)
8de758143786 stricter first-order check for meson
paulson
parents: 16563
diff changeset
   287
            in
8de758143786 stricter first-order check for meson
paulson
parents: 16563
diff changeset
   288
                if ok then
8de758143786 stricter first-order check for meson
paulson
parents: 16563
diff changeset
   289
                    let val (skoths,thy') = skolem thy (name, nnfth)
8de758143786 stricter first-order check for meson
paulson
parents: 16563
diff changeset
   290
			val cls = Meson.make_cnf skoths nnfth
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17404
diff changeset
   291
		    in change clause_cache (Symtab.update (name, (th, cls)));
16588
8de758143786 stricter first-order check for meson
paulson
parents: 16563
diff changeset
   292
			skolemlist nths thy'
8de758143786 stricter first-order check for meson
paulson
parents: 16563
diff changeset
   293
		    end
8de758143786 stricter first-order check for meson
paulson
parents: 16563
diff changeset
   294
		else skolemlist nths thy
8de758143786 stricter first-order check for meson
paulson
parents: 16563
diff changeset
   295
            end
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   296
	| SOME _ => skolemlist nths thy) (*FIXME: check for duplicate names?*)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   297
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   298
(*Exported function to convert Isabelle theorems into axiom clauses*) 
18000
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   299
fun cnf_axiom_g cnf_axiom_aux (name,th) =
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   300
    case name of
15955
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   301
	  "" => cnf_axiom_aux th (*no name, so can't cache*)
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17404
diff changeset
   302
	| s  => case Symtab.lookup (!clause_cache) s of
15955
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   303
	  	  NONE => 
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   304
		    let val cls = cnf_axiom_aux th
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17404
diff changeset
   305
		    in change clause_cache (Symtab.update (s, (th, cls))); cls end
15955
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   306
	        | SOME(th',cls) =>
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   307
		    if eq_thm(th,th') then cls
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   308
		    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
   309
		      let val cls = cnf_axiom_aux th
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17404
diff changeset
   310
		      in change clause_cache (Symtab.update (s, (th, cls))); cls end;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   311
18000
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   312
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   313
val cnf_axiom = cnf_axiom_g cnf_axiom_aux;
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   314
val cnf_axiomH = cnf_axiom_g cnf_axiom_auxH;
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   315
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   316
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   317
fun pairname th = (Thm.name_of_thm th, th);
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   318
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   319
fun meta_cnf_axiom th = 
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   320
    map Meson.make_meta_clause (cnf_axiom (pairname th));
15499
419dc5ffe8bc clausification and proof reconstruction
paulson
parents: 15495
diff changeset
   321
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   322
18000
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   323
fun meta_cnf_axiomH th = 
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   324
    map Meson.make_meta_clause (cnf_axiomH (pairname th));
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   325
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   326
(* changed: with one extra case added *)
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   327
fun univ_vars_of_aux (Const ("Hilbert_Choice.Eps",_) $ Abs(_,_,body)) vars =    
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   328
      univ_vars_of_aux body vars
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   329
  | univ_vars_of_aux (Const ("Ex",_) $ Abs(_,_,body)) vars = 
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   330
      univ_vars_of_aux body vars (* EX x. body *)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   331
  | univ_vars_of_aux (P $ Q) vars =
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   332
      univ_vars_of_aux Q (univ_vars_of_aux P vars)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   333
  | univ_vars_of_aux (t as Var(_,_)) vars = 
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   334
      if (t mem vars) then vars else (t::vars)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   335
  | univ_vars_of_aux _ vars = vars;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   336
  
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   337
fun univ_vars_of t = univ_vars_of_aux t [];
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   338
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   339
17819
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17484
diff changeset
   340
(**** Generating Skoklem Functions ****)
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17484
diff changeset
   341
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17484
diff changeset
   342
val skolem_prefix = "sk_";
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17484
diff changeset
   343
val skolem_id = ref 0;
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17484
diff changeset
   344
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17484
diff changeset
   345
fun gen_skolem_name () =
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17484
diff changeset
   346
  let val sk_fun = skolem_prefix ^ string_of_int (! skolem_id);
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17484
diff changeset
   347
  in inc skolem_id; sk_fun end;
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17484
diff changeset
   348
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17484
diff changeset
   349
fun gen_skolem univ_vars tp =
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17484
diff changeset
   350
  let
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17484
diff changeset
   351
    val skolem_type = map Term.fastype_of univ_vars ---> tp
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17484
diff changeset
   352
    val skolem_term = Const (gen_skolem_name (), skolem_type)
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17484
diff changeset
   353
  in Term.list_comb (skolem_term, univ_vars) end;
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17484
diff changeset
   354
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   355
fun get_new_skolem epss (t as (Const ("Hilbert_Choice.Eps",_) $ Abs(_,tp,_)))  = 
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   356
    let val all_vars = univ_vars_of t
17819
1241e5d31d5b small tidy-up of utility functions
paulson
parents: 17484
diff changeset
   357
	val sk_term = gen_skolem all_vars tp
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   358
    in
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   359
	(sk_term,(t,sk_term)::epss)
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   360
    end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   361
17404
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   362
(*FIXME: use a-list lookup!!*)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15499
diff changeset
   363
fun sk_lookup [] t = NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15499
diff changeset
   364
  | 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
   365
15390
87f78411f7c9 Comments and other tweaks by Jia
paulson
parents: 15371
diff changeset
   366
(* get the proper skolem term to replace epsilon term *)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   367
fun get_skolem epss t = 
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   368
    case (sk_lookup epss t) of NONE => get_new_skolem epss t
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   369
		             | SOME sk => (sk,epss);
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   370
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   371
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
   372
       get_skolem epss t
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   373
  | rm_Eps_cls_aux epss (P $ Q) =
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   374
       let val (P',epss') = rm_Eps_cls_aux epss P
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   375
	   val (Q',epss'') = rm_Eps_cls_aux epss' Q
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   376
       in (P' $ Q',epss'') end
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   377
  | rm_Eps_cls_aux epss t = (t,epss);
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   378
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   379
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
   380
17404
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   381
(* replace the epsilon terms in a formula by skolem terms. *)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   382
fun rm_Eps _ [] = []
17829
35123f89801e simplifying the treatment of clausification
paulson
parents: 17819
diff changeset
   383
  | rm_Eps epss (th::ths) = 
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   384
      let val (th',epss') = rm_Eps_cls epss th
17829
35123f89801e simplifying the treatment of clausification
paulson
parents: 17819
diff changeset
   385
      in th' :: (rm_Eps epss' ths) end;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   386
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   387
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   388
15872
8336ff711d80 fixed treatment of higher-order simprules
paulson
parents: 15736
diff changeset
   389
(**** Extract and Clausify theorems from a theory's claset and simpset ****)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   390
17404
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   391
(*Preserve the name of "th" after the transformation "f"*)
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   392
fun preserve_name f th = Thm.name_thm (Thm.name_of_thm th, f th);
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   393
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   394
(*Tags identify the major premise or conclusion, as hints to resolution provers.
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   395
  However, they don't appear to help in recent tests, and they complicate the code.*)
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   396
val tagI = thm "tagI";
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   397
val tagD = thm "tagD";
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   398
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   399
val tag_intro = preserve_name (fn th => th RS tagI);
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   400
val tag_elim  = preserve_name (fn th => tagD RS th);
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   401
17484
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17412
diff changeset
   402
fun rules_of_claset cs =
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17412
diff changeset
   403
  let val {safeIs,safeEs,hazIs,hazEs,...} = rep_cs cs
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17412
diff changeset
   404
      val intros = safeIs @ hazIs
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17412
diff changeset
   405
      val elims  = safeEs @ hazEs
17404
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   406
  in
17484
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17412
diff changeset
   407
     debug ("rules_of_claset intros: " ^ Int.toString(length intros) ^ 
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17412
diff changeset
   408
            " elims: " ^ Int.toString(length elims));
17404
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   409
     if tagging_enabled 
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   410
     then map pairname (map tag_intro intros @ map tag_elim elims)
17484
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17412
diff changeset
   411
     else map pairname (intros @ elims)
17404
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   412
  end;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   413
17484
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17412
diff changeset
   414
fun rules_of_simpset ss =
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17412
diff changeset
   415
  let val ({rules,...}, _) = rep_ss ss
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17412
diff changeset
   416
      val simps = Net.entries rules
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17412
diff changeset
   417
  in 
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17412
diff changeset
   418
      debug ("rules_of_simpset: " ^ Int.toString(length simps));
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17412
diff changeset
   419
      map (fn r => (#name r, #thm r)) simps
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17412
diff changeset
   420
  end;
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17412
diff changeset
   421
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17412
diff changeset
   422
fun claset_rules_of_thy thy = rules_of_claset (claset_of thy);
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17412
diff changeset
   423
fun simpset_rules_of_thy thy = rules_of_simpset (simpset_of thy);
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17412
diff changeset
   424
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17412
diff changeset
   425
fun claset_rules_of_ctxt ctxt = rules_of_claset (local_claset_of ctxt);
f6a225f97f0a simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents: 17412
diff changeset
   426
fun simpset_rules_of_ctxt ctxt = rules_of_simpset (local_simpset_of ctxt);
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   427
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   428
15872
8336ff711d80 fixed treatment of higher-order simprules
paulson
parents: 15736
diff changeset
   429
(**** 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
   430
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   431
(* classical rules *)
18000
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   432
fun cnf_rules_g cnf_axiom [] err_list = ([],err_list)
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   433
  | cnf_rules_g cnf_axiom ((name,th) :: ths) err_list = 
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   434
      let val (ts,es) = cnf_rules_g cnf_axiom ths err_list
17404
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   435
      in  (cnf_axiom (name,th) :: ts,es) handle  _ => (ts, (th::es))  end;  
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   436
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   437
18000
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   438
val cnf_rules = cnf_rules_g cnf_axiom;
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   439
val cnf_rulesH = cnf_rules_g cnf_axiomH;
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   440
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   441
15872
8336ff711d80 fixed treatment of higher-order simprules
paulson
parents: 15736
diff changeset
   442
(**** Convert all theorems of a claset/simpset into clauses (ResClause.clause) ****)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   443
17404
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   444
fun addclause (c,th) l =
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   445
  if ResClause.isTaut c then l else (c,th) :: l;
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   446
18000
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   447
fun addclauseH (c,th) l =
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   448
    if ResHolClause.isTaut c then l else (c,th)::l;
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   449
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   450
17829
35123f89801e simplifying the treatment of clausification
paulson
parents: 17819
diff changeset
   451
(* outputs a list of (clause,theorem) pairs *)
35123f89801e simplifying the treatment of clausification
paulson
parents: 17819
diff changeset
   452
fun clausify_axiom_pairs (thm_name,th) =
35123f89801e simplifying the treatment of clausification
paulson
parents: 17819
diff changeset
   453
    let val isa_clauses = cnf_axiom (thm_name,th) 
16039
dfe264950511 Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents: 16012
diff changeset
   454
        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
   455
        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
   456
	fun make_axiom_clauses _ [] []= []
16897
7e5319d0f418 code streamlining
paulson
parents: 16800
diff changeset
   457
	  | make_axiom_clauses i (cls::clss) (cls'::clss') =
17404
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   458
	      addclause (ResClause.make_axiom_clause cls (thm_name,i), cls') 
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   459
	                (make_axiom_clauses (i+1) clss clss')
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   460
    in
16039
dfe264950511 Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents: 16012
diff changeset
   461
	make_axiom_clauses 0 isa_clauses' isa_clauses		
17404
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   462
    end
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   463
18000
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   464
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   465
fun clausify_axiom_pairsH (thm_name,th) =
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   466
    let val isa_clauses = cnf_axiomH (thm_name,th) 
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   467
        val isa_clauses' = rm_Eps [] isa_clauses
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   468
        val clauses_n = length isa_clauses
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   469
	fun make_axiom_clauses _ [] []= []
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   470
	  | make_axiom_clauses i (cls::clss) (cls'::clss') =
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   471
	      addclauseH (ResHolClause.make_axiom_clause cls (thm_name,i), cls') 
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   472
	                (make_axiom_clauses (i+1) clss clss')
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   473
    in
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   474
	make_axiom_clauses 0 isa_clauses' isa_clauses		
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   475
    end
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   476
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   477
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   478
17829
35123f89801e simplifying the treatment of clausification
paulson
parents: 17819
diff changeset
   479
fun clausify_rules_pairs_aux result [] = result
35123f89801e simplifying the treatment of clausification
paulson
parents: 17819
diff changeset
   480
  | clausify_rules_pairs_aux result ((name,th)::ths) =
35123f89801e simplifying the treatment of clausification
paulson
parents: 17819
diff changeset
   481
      clausify_rules_pairs_aux (clausify_axiom_pairs (name,th) :: result) ths
35123f89801e simplifying the treatment of clausification
paulson
parents: 17819
diff changeset
   482
      handle THM (msg,_,_) =>  
35123f89801e simplifying the treatment of clausification
paulson
parents: 17819
diff changeset
   483
	      (debug ("Cannot clausify " ^ name ^ ": " ^ msg); 
35123f89801e simplifying the treatment of clausification
paulson
parents: 17819
diff changeset
   484
	       clausify_rules_pairs_aux result ths)
35123f89801e simplifying the treatment of clausification
paulson
parents: 17819
diff changeset
   485
	 |  ResClause.CLAUSE (msg,t) => 
35123f89801e simplifying the treatment of clausification
paulson
parents: 17819
diff changeset
   486
	      (debug ("Cannot clausify " ^ name ^ ": " ^ msg ^
35123f89801e simplifying the treatment of clausification
paulson
parents: 17819
diff changeset
   487
		      ": " ^ TermLib.string_of_term t); 
35123f89801e simplifying the treatment of clausification
paulson
parents: 17819
diff changeset
   488
	       clausify_rules_pairs_aux result ths)
17404
d16c3a62c396 the experimental tagging system, and the usual tidying
paulson
parents: 17279
diff changeset
   489
18000
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   490
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   491
fun clausify_rules_pairs_auxH result [] = result
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   492
  | clausify_rules_pairs_auxH result ((name,th)::ths) =
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   493
      clausify_rules_pairs_auxH (clausify_axiom_pairsH (name,th) :: result) ths
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   494
      handle THM (msg,_,_) =>  
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   495
	      (debug ("Cannot clausify " ^ name ^ ": " ^ msg); 
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   496
	       clausify_rules_pairs_auxH result ths)
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   497
	 |  ResHolClause.LAM2COMB (t) => 
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   498
	      (debug ("Cannot clausify "  ^ TermLib.string_of_term t); 
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   499
	       clausify_rules_pairs_auxH result ths)
16039
dfe264950511 Moved some of the clausify functions from ATP/res_clasimpset.ML to res_axioms.ML.
quigley
parents: 16012
diff changeset
   500
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   501
18000
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   502
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   503
val clausify_rules_pairs = clausify_rules_pairs_aux [];
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   504
ac059afd6b86 Added several new functions that convert HOL Isabelle rules to FOL axiom clauses. The original functions that convert FOL rules to clauses stay with the same names; the new functions have "H" at the end of their names.
mengj
parents: 17959
diff changeset
   505
val clausify_rules_pairsH = clausify_rules_pairs_auxH [];
16009
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   506
(*Setup function: takes a theory and installs ALL simprules and claset rules 
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   507
  into the clause cache*)
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   508
fun clause_cache_setup thy =
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   509
  let val simps = simpset_rules_of_thy thy
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   510
      and clas  = claset_rules_of_thy thy
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   511
  in skolemlist clas (skolemlist simps thy) end;
a6d480e6c5f0 Skolemization of simprules and classical rules
paulson
parents: 15997
diff changeset
   512
  
16563
a92f96951355 meson method taking an argument list
paulson
parents: 16173
diff changeset
   513
val clause_setup = [clause_cache_setup];  
a92f96951355 meson method taking an argument list
paulson
parents: 16173
diff changeset
   514
a92f96951355 meson method taking an argument list
paulson
parents: 16173
diff changeset
   515
a92f96951355 meson method taking an argument list
paulson
parents: 16173
diff changeset
   516
(*** meson proof methods ***)
a92f96951355 meson method taking an argument list
paulson
parents: 16173
diff changeset
   517
a92f96951355 meson method taking an argument list
paulson
parents: 16173
diff changeset
   518
fun cnf_rules_of_ths ths = List.concat (#1 (cnf_rules (map pairname ths) []));
a92f96951355 meson method taking an argument list
paulson
parents: 16173
diff changeset
   519
a92f96951355 meson method taking an argument list
paulson
parents: 16173
diff changeset
   520
fun meson_meth ths ctxt =
a92f96951355 meson method taking an argument list
paulson
parents: 16173
diff changeset
   521
  Method.SIMPLE_METHOD' HEADGOAL
a92f96951355 meson method taking an argument list
paulson
parents: 16173
diff changeset
   522
    (CHANGED_PROP o Meson.meson_claset_tac (cnf_rules_of_ths ths) (local_claset_of ctxt));
a92f96951355 meson method taking an argument list
paulson
parents: 16173
diff changeset
   523
a92f96951355 meson method taking an argument list
paulson
parents: 16173
diff changeset
   524
val meson_method_setup =
a92f96951355 meson method taking an argument list
paulson
parents: 16173
diff changeset
   525
 [Method.add_methods
a92f96951355 meson method taking an argument list
paulson
parents: 16173
diff changeset
   526
  [("meson", Method.thms_ctxt_args meson_meth, 
a92f96951355 meson method taking an argument list
paulson
parents: 16173
diff changeset
   527
    "The MESON resolution proof procedure")]];
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   528
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   529
end;