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