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