src/HOL/Tools/res_axioms.ML
author paulson
Thu, 12 May 2005 18:24:42 +0200
changeset 15956 0da64b5a9a00
parent 15955 87cf2ce8ede8
child 15997 c71031d7988c
permissions -rw-r--r--
theorem names for caching
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
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
     8
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
     9
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    10
signature RES_ELIM_RULE =
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    11
sig
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    12
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    13
exception ELIMR2FOL of string
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    14
val elimRule_tac : thm -> Tactical.tactic
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    15
val elimR2Fol : thm -> Term.term
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    16
val transform_elim : thm -> thm
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    17
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    18
end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    19
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    20
structure ResElimRule: RES_ELIM_RULE =
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    21
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    22
struct
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    23
15390
87f78411f7c9 Comments and other tweaks by Jia
paulson
parents: 15371
diff changeset
    24
(* a tactic used to prove an elim-rule. *)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    25
fun elimRule_tac thm =
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    26
    ((rtac impI 1) ORELSE (rtac notI 1)) THEN (etac thm 1) THEN
15371
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    27
    REPEAT(Fast_tac 1);
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    28
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    29
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    30
(* This following version fails sometimes, need to investigate, do not use it now. *)
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    31
fun elimRule_tac' thm =
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    32
   ((rtac impI 1) ORELSE (rtac notI 1)) THEN (etac thm 1) THEN
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    33
   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
    34
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    35
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    36
exception ELIMR2FOL of string;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    37
15390
87f78411f7c9 Comments and other tweaks by Jia
paulson
parents: 15371
diff changeset
    38
(* functions used to construct a formula *)
87f78411f7c9 Comments and other tweaks by Jia
paulson
parents: 15371
diff changeset
    39
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    40
fun make_disjs [x] = x
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    41
  | make_disjs (x :: xs) = HOLogic.mk_disj(x, make_disjs xs)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    42
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    43
fun make_conjs [x] = x
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    44
  | make_conjs (x :: xs) =  HOLogic.mk_conj(x, make_conjs xs)
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    45
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    46
fun add_EX tm [] = tm
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    47
  | 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
    48
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    49
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    50
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    51
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
    52
  | is_neg _ _ = false;
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    53
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    54
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    55
exception STRIP_CONCL;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    56
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    57
15371
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    58
fun strip_concl' prems bvs (Const ("==>",_) $ P $ Q) =
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    59
      let val P' = HOLogic.dest_Trueprop P
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    60
  	  val prems' = P'::prems
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    61
      in
15371
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    62
	strip_concl' prems' bvs  Q
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    63
      end
15371
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    64
  | strip_concl' prems bvs P = 
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    65
      let val P' = HOLogic.Not $ (HOLogic.dest_Trueprop P)
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    66
      in
15371
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    67
	add_EX (make_conjs (P'::prems)) bvs
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    68
      end;
15371
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    69
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    70
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    71
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
    72
  | strip_concl prems bvs concl (Const ("==>",_) $ P $ Q) =
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    73
    if (is_neg P concl) then (strip_concl' prems bvs Q)
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    74
    else
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    75
	(let val P' = HOLogic.dest_Trueprop P
15371
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    76
	     val prems' = P'::prems
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    77
	 in
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    78
	     strip_concl prems' bvs  concl Q
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    79
	 end)
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    80
  | strip_concl prems bvs concl _ = add_EX (make_conjs prems) bvs;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    81
 
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    82
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    83
15371
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    84
fun trans_elim (main,others,concl) =
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    85
    let val others' = map (strip_concl [] [] concl) others
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    86
	val disjs = make_disjs others'
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    87
    in
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    88
	HOLogic.mk_imp (HOLogic.dest_Trueprop main, disjs)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    89
    end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    90
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    91
15390
87f78411f7c9 Comments and other tweaks by Jia
paulson
parents: 15371
diff changeset
    92
(* aux function of elim2Fol, take away predicate variable. *)
15371
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    93
fun elimR2Fol_aux prems concl = 
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    94
    let val nprems = length prems
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    95
	val main = hd prems
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    96
    in
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
    97
	if (nprems = 1) then HOLogic.Not $ (HOLogic.dest_Trueprop main)
15371
40f5045c5985 fixes to clause conversion
paulson
parents: 15370
diff changeset
    98
        else trans_elim (main, tl prems, concl)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
    99
    end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   100
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   101
    
15390
87f78411f7c9 Comments and other tweaks by Jia
paulson
parents: 15371
diff changeset
   102
(* convert an elim rule into an equivalent formula, of type Term.term. *)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   103
fun elimR2Fol elimR = 
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   104
    let val elimR' = Drule.freeze_all elimR
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   105
	val (prems,concl) = (prems_of elimR', concl_of elimR')
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   106
    in
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   107
	case concl of Const("Trueprop",_) $ Free(_,Type("bool",[])) 
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   108
		      => HOLogic.mk_Trueprop (elimR2Fol_aux prems concl)
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   109
                    | Free(x,Type("prop",[])) => HOLogic.mk_Trueprop(elimR2Fol_aux prems concl) 
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   110
		    | _ => raise ELIMR2FOL("Not an elimination rule!")
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   111
    end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   112
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   113
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   114
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   115
(**** use prove_goalw_cterm to prove ****)
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   116
15390
87f78411f7c9 Comments and other tweaks by Jia
paulson
parents: 15371
diff changeset
   117
(* convert an elim-rule into an equivalent theorem that does not have the predicate variable. *) 
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   118
fun transform_elim thm =
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   119
    let val tm = elimR2Fol thm
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   120
	val ctm = cterm_of (sign_of_thm thm) tm	
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   121
    in
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   122
	prove_goalw_cterm [] ctm (fn prems => [elimRule_tac thm])
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   123
    end;	
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   124
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   125
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   126
end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   127
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   128
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   129
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   130
signature RES_AXIOMS =
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   131
sig
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   132
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   133
val clausify_axiom : thm -> ResClause.clause list
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   134
val cnf_axiom : (string * thm) -> thm list
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   135
val meta_cnf_axiom : thm -> thm list
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   136
val cnf_elim : thm -> thm list
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   137
val cnf_rule : thm -> thm list
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   138
val cnf_classical_rules_thy : theory -> thm list list * thm list
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   139
val clausify_classical_rules_thy : theory -> ResClause.clause list list * thm list
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   140
val cnf_simpset_rules_thy : theory -> thm list list * thm list
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   141
val clausify_simpset_rules_thy : theory -> ResClause.clause list list * thm list
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   142
val rm_Eps 
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   143
: (Term.term * Term.term) list -> thm list -> Term.term list
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   144
val claset_rules_of_thy : theory -> (string * thm) list
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   145
val simpset_rules_of_thy : theory -> (string * thm) list
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   146
val clausify_rules : thm list -> thm list -> ResClause.clause list list * thm list
15684
5ec4d21889d6 Reconstruction code, now packaged to avoid name clashes
paulson
parents: 15644
diff changeset
   147
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   148
end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   149
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   150
structure ResAxioms : RES_AXIOMS =
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   151
 
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   152
struct
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   153
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   154
open ResElimRule;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   155
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   156
(* to be fixed: cnf_intro, cnf_rule, is_introR *)
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   157
15390
87f78411f7c9 Comments and other tweaks by Jia
paulson
parents: 15371
diff changeset
   158
(* check if a rule is an elim rule *)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   159
fun is_elimR thm = 
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   160
    case (concl_of thm) of (Const ("Trueprop", _) $ Var (idx,_)) => true
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   161
			 | Var(indx,Type("prop",[])) => true
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   162
			 | _ => false;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   163
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   164
15390
87f78411f7c9 Comments and other tweaks by Jia
paulson
parents: 15371
diff changeset
   165
(* repeated resolution *)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   166
fun repeat_RS thm1 thm2 =
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   167
    let val thm1' =  thm1 RS thm2 handle THM _ => thm1
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   168
    in
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   169
	if eq_thm(thm1,thm1') then thm1' else (repeat_RS thm1' thm2)
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   170
    end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   171
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   172
15390
87f78411f7c9 Comments and other tweaks by Jia
paulson
parents: 15371
diff changeset
   173
(* convert a theorem into NNF and also skolemize it. *)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   174
fun skolem_axiom thm = 
15872
8336ff711d80 fixed treatment of higher-order simprules
paulson
parents: 15736
diff changeset
   175
  if Term.is_first_order (prop_of thm) then
8336ff711d80 fixed treatment of higher-order simprules
paulson
parents: 15736
diff changeset
   176
    let val thm' = (skolemize o make_nnf o ObjectLogic.atomize_thm o Drule.freeze_all) thm
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   177
    in 
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   178
	repeat_RS thm' someI_ex
15872
8336ff711d80 fixed treatment of higher-order simprules
paulson
parents: 15736
diff changeset
   179
    end
8336ff711d80 fixed treatment of higher-order simprules
paulson
parents: 15736
diff changeset
   180
  else raise THM ("skolem_axiom: not first-order", 0, [thm]);
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   181
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   182
15872
8336ff711d80 fixed treatment of higher-order simprules
paulson
parents: 15736
diff changeset
   183
fun cnf_rule thm = make_clauses [skolem_axiom thm]
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   184
15872
8336ff711d80 fixed treatment of higher-order simprules
paulson
parents: 15736
diff changeset
   185
fun cnf_elim thm = cnf_rule (transform_elim thm);
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   186
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   187
15370
05b03ea0f18d trying to fix the transfer problem
paulson
parents: 15359
diff changeset
   188
(*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
   189
  inside that theory -- because it's needed for Skolemization *)
8bad1f42fec0 new CLAUSIFY attribute for proof reconstruction with lemmas
paulson
parents: 15347
diff changeset
   190
15370
05b03ea0f18d trying to fix the transfer problem
paulson
parents: 15359
diff changeset
   191
val recon_thy = ThyInfo.get_theory"Reconstruction";
15359
8bad1f42fec0 new CLAUSIFY attribute for proof reconstruction with lemmas
paulson
parents: 15347
diff changeset
   192
15370
05b03ea0f18d trying to fix the transfer problem
paulson
parents: 15359
diff changeset
   193
fun transfer_to_Reconstruction thm =
05b03ea0f18d trying to fix the transfer problem
paulson
parents: 15359
diff changeset
   194
    transfer recon_thy thm handle THM _ => thm;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   195
15955
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   196
fun is_taut th =
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   197
      case (prop_of th) of
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   198
           (Const ("Trueprop", _) $ Const ("True", _)) => true
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   199
         | _ => false;
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   200
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   201
(* remove tautologous clauses *)
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   202
val rm_redundant_cls = List.filter (not o is_taut);
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   203
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   204
(* transform an Isabelle thm into CNF *)
15955
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   205
fun cnf_axiom_aux thm =
15370
05b03ea0f18d trying to fix the transfer problem
paulson
parents: 15359
diff changeset
   206
    let val thm' = transfer_to_Reconstruction thm
15499
419dc5ffe8bc clausification and proof reconstruction
paulson
parents: 15495
diff changeset
   207
	val thm'' = if (is_elimR thm') then (cnf_elim thm')  else cnf_rule thm'
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   208
    in
15955
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   209
	map (zero_var_indexes o Thm.varifyT) (rm_redundant_cls thm'')
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   210
    end;
15955
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   211
    
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   212
(*Cache for clauses: could be a hash table if we provided them.*)
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   213
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
   214
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   215
fun cnf_axiom (name,th) =
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   216
    case name of
15955
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   217
	  "" => 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
   218
	| 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
   219
	  	  NONE => 
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   220
		    let val cls = cnf_axiom_aux th
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   221
		    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
   222
		    end
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   223
	        | SOME(th',cls) =>
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   224
		    if eq_thm(th,th') then cls
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   225
		    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
   226
		      let val cls = cnf_axiom_aux th
87cf2ce8ede8 memoization of ResAxioms.cnf_axiom rather than of Reconstruction.clausify_rule
paulson
parents: 15872
diff changeset
   227
		      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
   228
		      end;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   229
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   230
fun pairname th = (Thm.name_of_thm th, th);
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   231
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   232
fun meta_cnf_axiom th = 
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   233
    map Meson.make_meta_clause (cnf_axiom (pairname th));
15499
419dc5ffe8bc clausification and proof reconstruction
paulson
parents: 15495
diff changeset
   234
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   235
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   236
(* changed: with one extra case added *)
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   237
fun univ_vars_of_aux (Const ("Hilbert_Choice.Eps",_) $ Abs(_,_,body)) vars =    
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   238
      univ_vars_of_aux body vars
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   239
  | univ_vars_of_aux (Const ("Ex",_) $ Abs(_,_,body)) vars = 
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   240
      univ_vars_of_aux body vars (* EX x. body *)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   241
  | univ_vars_of_aux (P $ Q) vars =
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   242
      univ_vars_of_aux Q (univ_vars_of_aux P vars)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   243
  | univ_vars_of_aux (t as Var(_,_)) vars = 
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   244
      if (t mem vars) then vars else (t::vars)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   245
  | univ_vars_of_aux _ vars = vars;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   246
  
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   247
fun univ_vars_of t = univ_vars_of_aux t [];
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   248
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   249
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   250
fun get_new_skolem epss (t as (Const ("Hilbert_Choice.Eps",_) $ Abs(_,tp,_)))  = 
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   251
    let val all_vars = univ_vars_of t
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   252
	val sk_term = ResSkolemFunction.gen_skolem all_vars tp
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   253
    in
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   254
	(sk_term,(t,sk_term)::epss)
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   255
    end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   256
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   257
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15499
diff changeset
   258
fun sk_lookup [] t = NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15499
diff changeset
   259
  | 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
   260
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   261
15390
87f78411f7c9 Comments and other tweaks by Jia
paulson
parents: 15371
diff changeset
   262
87f78411f7c9 Comments and other tweaks by Jia
paulson
parents: 15371
diff changeset
   263
(* get the proper skolem term to replace epsilon term *)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   264
fun get_skolem epss t = 
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   265
    case (sk_lookup epss t) of NONE => get_new_skolem epss t
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   266
		             | SOME sk => (sk,epss);
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   267
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   268
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   269
fun rm_Eps_cls_aux epss (t as (Const ("Hilbert_Choice.Eps",_) $ Abs(_,_,_))) = get_skolem epss t
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   270
  | rm_Eps_cls_aux epss (P $ Q) =
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   271
    let val (P',epss') = rm_Eps_cls_aux epss P
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   272
	val (Q',epss'') = rm_Eps_cls_aux epss' Q
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   273
    in
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   274
	(P' $ Q',epss'')
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   275
    end
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   276
  | rm_Eps_cls_aux epss t = (t,epss);
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   277
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   278
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   279
fun rm_Eps_cls epss thm = rm_Eps_cls_aux epss (prop_of thm);
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   280
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   281
15390
87f78411f7c9 Comments and other tweaks by Jia
paulson
parents: 15371
diff changeset
   282
(* remove the epsilon terms in a formula, by skolem terms. *)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   283
fun rm_Eps _ [] = []
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   284
  | rm_Eps epss (thm::thms) = 
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   285
      let val (thm',epss') = rm_Eps_cls epss thm
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   286
      in
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   287
	thm' :: (rm_Eps epss' thms)
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   288
      end;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   289
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   290
15390
87f78411f7c9 Comments and other tweaks by Jia
paulson
parents: 15371
diff changeset
   291
(* convert a theorem into CNF and then into Clause.clause format. *)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   292
fun clausify_axiom thm =
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   293
    let val name = Thm.name_of_thm thm
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   294
	val isa_clauses = cnf_axiom (name, thm)
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   295
	      (*"isa_clauses" are already "standard"ed. *)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   296
        val isa_clauses' = rm_Eps [] isa_clauses
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   297
        val clauses_n = length isa_clauses
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   298
	fun make_axiom_clauses _ [] = []
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   299
	  | make_axiom_clauses i (cls::clss) = (ResClause.make_axiom_clause cls (name,i)) :: make_axiom_clauses (i+1) clss 
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   300
    in
15872
8336ff711d80 fixed treatment of higher-order simprules
paulson
parents: 15736
diff changeset
   301
	make_axiom_clauses 0 isa_clauses'		
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   302
    end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   303
  
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   304
15872
8336ff711d80 fixed treatment of higher-order simprules
paulson
parents: 15736
diff changeset
   305
(**** Extract and Clausify theorems from a theory's claset and simpset ****)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   306
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   307
fun claset_rules_of_thy thy =
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   308
    let val clsset = rep_cs (claset_of thy)
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   309
	val safeEs = #safeEs clsset
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   310
	val safeIs = #safeIs clsset
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   311
	val hazEs = #hazEs clsset
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   312
	val hazIs = #hazIs clsset
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   313
    in
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   314
	map pairname (safeEs @ safeIs @ hazEs @ hazIs)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   315
    end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   316
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   317
fun simpset_rules_of_thy thy =
15872
8336ff711d80 fixed treatment of higher-order simprules
paulson
parents: 15736
diff changeset
   318
    let val rules = #rules(fst (rep_ss (simpset_of thy)))
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   319
    in
15872
8336ff711d80 fixed treatment of higher-order simprules
paulson
parents: 15736
diff changeset
   320
	map (fn (_,r) => (#name r, #thm r)) (Net.dest rules)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   321
    end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   322
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   323
15872
8336ff711d80 fixed treatment of higher-order simprules
paulson
parents: 15736
diff changeset
   324
(**** 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
   325
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   326
(* classical rules *)
15872
8336ff711d80 fixed treatment of higher-order simprules
paulson
parents: 15736
diff changeset
   327
fun cnf_rules [] err_list = ([],err_list)
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   328
  | cnf_rules ((name,thm) :: thms) err_list = 
15872
8336ff711d80 fixed treatment of higher-order simprules
paulson
parents: 15736
diff changeset
   329
      let val (ts,es) = cnf_rules thms err_list
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   330
      in  (cnf_axiom (name,thm) :: ts,es) handle  _ => (ts, (thm::es))  end;
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   331
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   332
(* CNF all rules from a given theory's classical reasoner *)
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   333
fun cnf_classical_rules_thy thy = 
15872
8336ff711d80 fixed treatment of higher-order simprules
paulson
parents: 15736
diff changeset
   334
    cnf_rules (claset_rules_of_thy thy) [];
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   335
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   336
(* CNF all simplifier rules from a given theory's simpset *)
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   337
fun cnf_simpset_rules_thy thy =
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   338
    cnf_rules (simpset_rules_of_thy thy) [];
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
15872
8336ff711d80 fixed treatment of higher-order simprules
paulson
parents: 15736
diff changeset
   341
(**** Convert all theorems of a claset/simpset into clauses (ResClause.clause) ****)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   342
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   343
(* classical rules *)
15872
8336ff711d80 fixed treatment of higher-order simprules
paulson
parents: 15736
diff changeset
   344
fun clausify_rules [] err_list = ([],err_list)
8336ff711d80 fixed treatment of higher-order simprules
paulson
parents: 15736
diff changeset
   345
  | clausify_rules (thm::thms) err_list =
8336ff711d80 fixed treatment of higher-order simprules
paulson
parents: 15736
diff changeset
   346
    let val (ts,es) = clausify_rules thms err_list
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   347
    in
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   348
	((clausify_axiom thm)::ts,es) handle  _ => (ts,(thm::es))
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   349
    end;
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   350
15390
87f78411f7c9 Comments and other tweaks by Jia
paulson
parents: 15371
diff changeset
   351
15736
1bb0399a9517 more tidying up of the SPASS interface
paulson
parents: 15684
diff changeset
   352
(* convert all classical rules from a given theory into Clause.clause format. *)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   353
fun clausify_classical_rules_thy thy =
15956
0da64b5a9a00 theorem names for caching
paulson
parents: 15955
diff changeset
   354
    clausify_rules (map #2 (claset_rules_of_thy thy)) [];
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   355
15736
1bb0399a9517 more tidying up of the SPASS interface
paulson
parents: 15684
diff changeset
   356
(* convert all simplifier rules from a given theory into Clause.clause format. *)
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   357
fun clausify_simpset_rules_thy thy =
15872
8336ff711d80 fixed treatment of higher-order simprules
paulson
parents: 15736
diff changeset
   358
    clausify_rules (map #2 (simpset_rules_of_thy thy)) [];
15347
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   359
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   360
14585bc8fa09 resolution package tools by Jia Meng
paulson
parents:
diff changeset
   361
end;