| author | haftmann | 
| Wed, 30 Jul 2008 07:34:00 +0200 | |
| changeset 27710 | 29702aa892a5 | 
| parent 27330 | 1af2598b5f7d | 
| child 27809 | a1e409db516b | 
| permissions | -rw-r--r-- | 
| 15347 | 1  | 
(* Author: Jia Meng, Cambridge University Computer Laboratory  | 
2  | 
ID: $Id$  | 
|
3  | 
Copyright 2004 University of Cambridge  | 
|
4  | 
||
| 
20461
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
5  | 
Transformation of axiom rules (elim/intro/etc) into CNF forms.  | 
| 15347 | 6  | 
*)  | 
7  | 
||
| 15997 | 8  | 
signature RES_AXIOMS =  | 
| 21505 | 9  | 
sig  | 
| 
27179
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
10  | 
val cnf_axiom: theory -> thm -> thm list  | 
| 24669 | 11  | 
val pairname: thm -> string * thm  | 
| 27184 | 12  | 
val multi_base_blacklist: string list  | 
| 25243 | 13  | 
val bad_for_atp: thm -> bool  | 
| 25761 | 14  | 
val type_has_empty_sort: typ -> bool  | 
| 
27179
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
15  | 
val cnf_rules_pairs: theory -> (string * thm) list -> (thm * (string * int)) list  | 
| 24669 | 16  | 
val neg_clausify: thm list -> thm list  | 
17  | 
val expand_defs_tac: thm -> tactic  | 
|
| 24827 | 18  | 
val combinators: thm -> thm  | 
| 
21999
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21900 
diff
changeset
 | 
19  | 
val neg_conjecture_clauses: thm -> int -> thm list * (string * typ) list  | 
| 
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21900 
diff
changeset
 | 
20  | 
val claset_rules_of: Proof.context -> (string * thm) list (*FIXME DELETE*)  | 
| 
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21900 
diff
changeset
 | 
21  | 
val simpset_rules_of: Proof.context -> (string * thm) list (*FIXME DELETE*)  | 
| 21505 | 22  | 
val atpset_rules_of: Proof.context -> (string * thm) list  | 
| 
25256
 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 
paulson 
parents: 
25243 
diff
changeset
 | 
23  | 
val suppress_endtheory: bool ref (*for emergency use where endtheory causes problems*)  | 
| 24669 | 24  | 
val setup: theory -> theory  | 
| 21505 | 25  | 
end;  | 
| 
19196
 
62ee8c10d796
Added functions to retrieve local and global atpset rules.
 
mengj 
parents: 
19175 
diff
changeset
 | 
26  | 
|
| 24669 | 27  | 
structure ResAxioms: RES_AXIOMS =  | 
| 15997 | 28  | 
struct  | 
| 15347 | 29  | 
|
| 20902 | 30  | 
(* FIXME legacy *)  | 
| 
20863
 
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
 
paulson 
parents: 
20789 
diff
changeset
 | 
31  | 
fun freeze_thm th = #1 (Drule.freeze_thaw th);  | 
| 
 
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
 
paulson 
parents: 
20789 
diff
changeset
 | 
32  | 
|
| 25761 | 33  | 
fun type_has_empty_sort (TFree (_, [])) = true  | 
34  | 
| type_has_empty_sort (TVar (_, [])) = true  | 
|
35  | 
| type_has_empty_sort (Type (_, Ts)) = exists type_has_empty_sort Ts  | 
|
36  | 
| type_has_empty_sort _ = false;  | 
|
| 27184 | 37  | 
|
| 15997 | 38  | 
(**** Transformation of Elimination Rules into First-Order Formulas****)  | 
| 15347 | 39  | 
|
| 
21430
 
77651b6d9d6c
New transformation of eliminatino rules: we simply replace the final conclusion variable by False
 
paulson 
parents: 
21290 
diff
changeset
 | 
40  | 
val cfalse = cterm_of HOL.thy HOLogic.false_const;  | 
| 
 
77651b6d9d6c
New transformation of eliminatino rules: we simply replace the final conclusion variable by False
 
paulson 
parents: 
21290 
diff
changeset
 | 
41  | 
val ctp_false = cterm_of HOL.thy (HOLogic.mk_Trueprop HOLogic.false_const);  | 
| 
20461
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
42  | 
|
| 
21430
 
77651b6d9d6c
New transformation of eliminatino rules: we simply replace the final conclusion variable by False
 
paulson 
parents: 
21290 
diff
changeset
 | 
43  | 
(*Converts an elim-rule into an equivalent theorem that does not have the  | 
| 
 
77651b6d9d6c
New transformation of eliminatino rules: we simply replace the final conclusion variable by False
 
paulson 
parents: 
21290 
diff
changeset
 | 
44  | 
predicate variable. Leaves other theorems unchanged. We simply instantiate the  | 
| 
 
77651b6d9d6c
New transformation of eliminatino rules: we simply replace the final conclusion variable by False
 
paulson 
parents: 
21290 
diff
changeset
 | 
45  | 
conclusion variable to False.*)  | 
| 16009 | 46  | 
fun transform_elim th =  | 
| 
21430
 
77651b6d9d6c
New transformation of eliminatino rules: we simply replace the final conclusion variable by False
 
paulson 
parents: 
21290 
diff
changeset
 | 
47  | 
case concl_of th of (*conclusion variable*)  | 
| 24669 | 48  | 
       Const("Trueprop",_) $ (v as Var(_,Type("bool",[]))) =>
 | 
| 
21430
 
77651b6d9d6c
New transformation of eliminatino rules: we simply replace the final conclusion variable by False
 
paulson 
parents: 
21290 
diff
changeset
 | 
49  | 
Thm.instantiate ([], [(cterm_of HOL.thy v, cfalse)]) th  | 
| 24669 | 50  | 
    | v as Var(_, Type("prop",[])) =>
 | 
| 
21430
 
77651b6d9d6c
New transformation of eliminatino rules: we simply replace the final conclusion variable by False
 
paulson 
parents: 
21290 
diff
changeset
 | 
51  | 
Thm.instantiate ([], [(cterm_of HOL.thy v, ctp_false)]) th  | 
| 
 
77651b6d9d6c
New transformation of eliminatino rules: we simply replace the final conclusion variable by False
 
paulson 
parents: 
21290 
diff
changeset
 | 
52  | 
| _ => th;  | 
| 15997 | 53  | 
|
| 
24742
 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 
paulson 
parents: 
24669 
diff
changeset
 | 
54  | 
(*To enforce single-threading*)  | 
| 
 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 
paulson 
parents: 
24669 
diff
changeset
 | 
55  | 
exception Clausify_failure of theory;  | 
| 
20461
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
56  | 
|
| 16009 | 57  | 
(**** SKOLEMIZATION BY INFERENCE (lcp) ****)  | 
58  | 
||
| 
24742
 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 
paulson 
parents: 
24669 
diff
changeset
 | 
59  | 
fun rhs_extra_types lhsT rhs =  | 
| 
 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 
paulson 
parents: 
24669 
diff
changeset
 | 
60  | 
let val lhs_vars = Term.add_tfreesT lhsT []  | 
| 
 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 
paulson 
parents: 
24669 
diff
changeset
 | 
61  | 
fun add_new_TFrees (TFree v) =  | 
| 
24821
 
cc55041ca8ba
skolem_cache: ignore internal theorems -- major speedup;
 
wenzelm 
parents: 
24785 
diff
changeset
 | 
62  | 
if member (op =) lhs_vars v then I else insert (op =) (TFree v)  | 
| 
 
cc55041ca8ba
skolem_cache: ignore internal theorems -- major speedup;
 
wenzelm 
parents: 
24785 
diff
changeset
 | 
63  | 
| add_new_TFrees _ = I  | 
| 
24742
 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 
paulson 
parents: 
24669 
diff
changeset
 | 
64  | 
val rhs_consts = fold_aterms (fn Const c => insert (op =) c | _ => I) rhs []  | 
| 
 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 
paulson 
parents: 
24669 
diff
changeset
 | 
65  | 
in fold (#2 #> Term.fold_atyps add_new_TFrees) rhs_consts [] end;  | 
| 
 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 
paulson 
parents: 
24669 
diff
changeset
 | 
66  | 
|
| 
18141
 
89e2e8bed08f
Skolemization by inference, but not quite finished
 
paulson 
parents: 
18009 
diff
changeset
 | 
67  | 
(*Traverse a theorem, declaring Skolem function definitions. String s is the suggested  | 
| 
27174
 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
68  | 
prefix for the Skolem constant.*)  | 
| 
 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
69  | 
fun declare_skofuns s th =  | 
| 
 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
70  | 
let  | 
| 
 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
71  | 
val nref = ref 0  | 
| 
 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
72  | 
    fun dec_sko (Const ("Ex",_) $ (xtp as Abs (_, T, p))) (axs, thy) =
 | 
| 
 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
73  | 
(*Existential: declare a Skolem function, then insert into body and continue*)  | 
| 
 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
74  | 
let  | 
| 
 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
75  | 
val cname = "sko_" ^ s ^ "_" ^ Int.toString (inc nref)  | 
| 
 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
76  | 
val args0 = term_frees xtp (*get the formal parameter list*)  | 
| 
 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
77  | 
val Ts = map type_of args0  | 
| 
 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
78  | 
val extraTs = rhs_extra_types (Ts ---> T) xtp  | 
| 
 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
79  | 
val _ = if null extraTs then () else  | 
| 
 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
80  | 
              warning ("Skolemization: extra type vars: " ^
 | 
| 
 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
81  | 
commas_quote (map (Syntax.string_of_typ_global thy) extraTs))  | 
| 
 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
82  | 
val argsx = map (fn T => Free (gensym "vsk", T)) extraTs  | 
| 
 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
83  | 
val args = argsx @ args0  | 
| 
 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
84  | 
val cT = extraTs ---> Ts ---> T  | 
| 
 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
85  | 
val rhs = list_abs_free (map dest_Free args, HOLogic.choice_const T $ xtp)  | 
| 
 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
86  | 
(*Forms a lambda-abstraction over the formal parameters*)  | 
| 
 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
87  | 
val (c, thy') = Sign.declare_const [Markup.property_internal] (cname, cT, NoSyn) thy  | 
| 
 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
88  | 
val cdef = cname ^ "_def"  | 
| 27330 | 89  | 
val thy'' = Theory.add_defs_i true false [(cdef, Logic.mk_equals (c, rhs))] thy'  | 
| 
27174
 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
90  | 
val ax = Thm.get_axiom_i thy'' (Sign.full_name thy'' cdef)  | 
| 
 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
91  | 
in dec_sko (subst_bound (list_comb (c, args), p)) (ax :: axs, thy'') end  | 
| 
 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
92  | 
      | dec_sko (Const ("All", _) $ (xtp as Abs (a, T, p))) thx =
 | 
| 
 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
93  | 
(*Universal quant: insert a free variable into body and continue*)  | 
| 
 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
94  | 
let val fname = Name.variant (add_term_names (p, [])) a  | 
| 
 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
95  | 
in dec_sko (subst_bound (Free (fname, T), p)) thx end  | 
| 
 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
96  | 
      | dec_sko (Const ("op &", _) $ p $ q) thx = dec_sko q (dec_sko p thx)
 | 
| 
 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
97  | 
      | dec_sko (Const ("op |", _) $ p $ q) thx = dec_sko q (dec_sko p thx)
 | 
| 
 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
98  | 
      | dec_sko (Const ("Trueprop", _) $ p) thx = dec_sko p thx
 | 
| 
 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
99  | 
| dec_sko t thx = thx (*Do nothing otherwise*)  | 
| 
 
c2c484480f40
declare_skofuns/skolem: canonical argument order;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
100  | 
in fn thy => dec_sko (Thm.prop_of th) ([], thy) end;  | 
| 
18141
 
89e2e8bed08f
Skolemization by inference, but not quite finished
 
paulson 
parents: 
18009 
diff
changeset
 | 
101  | 
|
| 
 
89e2e8bed08f
Skolemization by inference, but not quite finished
 
paulson 
parents: 
18009 
diff
changeset
 | 
102  | 
(*Traverse a theorem, accumulating Skolem function definitions.*)  | 
| 
22731
 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 
paulson 
parents: 
22724 
diff
changeset
 | 
103  | 
fun assume_skofuns s th =  | 
| 
 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 
paulson 
parents: 
22724 
diff
changeset
 | 
104  | 
let val sko_count = ref 0  | 
| 
 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 
paulson 
parents: 
22724 
diff
changeset
 | 
105  | 
      fun dec_sko (Const ("Ex",_) $ (xtp as Abs(_,T,p))) defs =
 | 
| 
20461
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
106  | 
(*Existential: declare a Skolem function, then insert into body and continue*)  | 
| 
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
107  | 
let val skos = map (#1 o Logic.dest_equals) defs (*existing sko fns*)  | 
| 
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
108  | 
val args = term_frees xtp \\ skos (*the formal parameters*)  | 
| 
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
109  | 
val Ts = map type_of args  | 
| 
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
110  | 
val cT = Ts ---> T  | 
| 
22731
 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 
paulson 
parents: 
22724 
diff
changeset
 | 
111  | 
val id = "sko_" ^ s ^ "_" ^ Int.toString (inc sko_count)  | 
| 
 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 
paulson 
parents: 
22724 
diff
changeset
 | 
112  | 
val c = Free (id, cT)  | 
| 
20461
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
113  | 
val rhs = list_abs_free (map dest_Free args,  | 
| 
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
114  | 
HOLogic.choice_const T $ xtp)  | 
| 
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
115  | 
(*Forms a lambda-abstraction over the formal parameters*)  | 
| 27330 | 116  | 
val def = Logic.mk_equals (c, rhs)  | 
| 
20461
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
117  | 
in dec_sko (subst_bound (list_comb(c,args), p))  | 
| 
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
118  | 
(def :: defs)  | 
| 
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
119  | 
end  | 
| 
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
120  | 
        | dec_sko (Const ("All",_) $ (xtp as Abs(a,T,p))) defs =
 | 
| 
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
121  | 
(*Universal quant: insert a free variable into body and continue*)  | 
| 
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
122  | 
let val fname = Name.variant (add_term_names (p,[])) a  | 
| 
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
123  | 
in dec_sko (subst_bound (Free(fname,T), p)) defs end  | 
| 
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
124  | 
        | dec_sko (Const ("op &", _) $ p $ q) defs = dec_sko q (dec_sko p defs)
 | 
| 
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
125  | 
        | dec_sko (Const ("op |", _) $ p $ q) defs = dec_sko q (dec_sko p defs)
 | 
| 
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
126  | 
        | dec_sko (Const ("Trueprop", _) $ p) defs = dec_sko p defs
 | 
| 
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
127  | 
| dec_sko t defs = defs (*Do nothing otherwise*)  | 
| 20419 | 128  | 
in dec_sko (prop_of th) [] end;  | 
129  | 
||
130  | 
||
| 24827 | 131  | 
(**** REPLACING ABSTRACTIONS BY COMBINATORS ****)  | 
| 20419 | 132  | 
|
133  | 
(*Returns the vars of a theorem*)  | 
|
134  | 
fun vars_of_thm th =  | 
|
| 22691 | 135  | 
map (Thm.cterm_of (theory_of_thm th) o Var) (Thm.fold_terms Term.add_vars th []);  | 
| 20419 | 136  | 
|
137  | 
(*Make a version of fun_cong with a given variable name*)  | 
|
138  | 
local  | 
|
139  | 
val fun_cong' = fun_cong RS asm_rl; (*renumber f, g to prevent clashes with (a,0)*)  | 
|
140  | 
val cx = hd (vars_of_thm fun_cong');  | 
|
141  | 
val ty = typ_of (ctyp_of_term cx);  | 
|
| 20445 | 142  | 
val thy = theory_of_thm fun_cong;  | 
| 20419 | 143  | 
fun mkvar a = cterm_of thy (Var((a,0),ty));  | 
144  | 
in  | 
|
145  | 
fun xfun_cong x = Thm.instantiate ([], [(cx, mkvar x)]) fun_cong'  | 
|
146  | 
end;  | 
|
147  | 
||
| 
20863
 
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
 
paulson 
parents: 
20789 
diff
changeset
 | 
148  | 
(*Removes the lambdas from an equation of the form t = (%x. u). A non-negative n,  | 
| 
 
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
 
paulson 
parents: 
20789 
diff
changeset
 | 
149  | 
serves as an upper bound on how many to remove.*)  | 
| 
 
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
 
paulson 
parents: 
20789 
diff
changeset
 | 
150  | 
fun strip_lambdas 0 th = th  | 
| 24669 | 151  | 
| strip_lambdas n th =  | 
| 
20863
 
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
 
paulson 
parents: 
20789 
diff
changeset
 | 
152  | 
case prop_of th of  | 
| 24669 | 153  | 
          _ $ (Const ("op =", _) $ _ $ Abs (x,_,_)) =>
 | 
154  | 
strip_lambdas (n-1) (freeze_thm (th RS xfun_cong x))  | 
|
155  | 
| _ => th;  | 
|
| 20419 | 156  | 
|
| 24669 | 157  | 
val lambda_free = not o Term.has_abs;  | 
| 
20461
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
158  | 
|
| 24669 | 159  | 
val monomorphic = not o Term.exists_type (Term.exists_subtype Term.is_TVar);  | 
| 
20461
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
160  | 
|
| 27184 | 161  | 
val [f_B,g_B] = map (cterm_of @{theory}) (term_vars (prop_of @{thm abs_B}));
 | 
162  | 
val [g_C,f_C] = map (cterm_of @{theory}) (term_vars (prop_of @{thm abs_C}));
 | 
|
163  | 
val [f_S,g_S] = map (cterm_of @{theory}) (term_vars (prop_of @{thm abs_S}));
 | 
|
| 
20863
 
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
 
paulson 
parents: 
20789 
diff
changeset
 | 
164  | 
|
| 24827 | 165  | 
(*FIXME: requires more use of cterm constructors*)  | 
166  | 
fun abstract ct =  | 
|
| 26928 | 167  | 
let val _ = Output.debug (fn()=>" abstraction: " ^ Display.string_of_cterm ct)  | 
| 
25256
 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 
paulson 
parents: 
25243 
diff
changeset
 | 
168  | 
val Abs(x,_,body) = term_of ct  | 
| 24827 | 169  | 
val thy = theory_of_cterm ct  | 
170  | 
      val Type("fun",[xT,bodyT]) = typ_of (ctyp_of_term ct)
 | 
|
171  | 
val cxT = ctyp_of thy xT and cbodyT = ctyp_of thy bodyT  | 
|
| 27184 | 172  | 
      fun makeK() = instantiate' [SOME cxT, SOME cbodyT] [SOME (cterm_of thy body)] @{thm abs_K}
 | 
| 24827 | 173  | 
in  | 
174  | 
case body of  | 
|
175  | 
Const _ => makeK()  | 
|
176  | 
| Free _ => makeK()  | 
|
177  | 
| Var _ => makeK() (*though Var isn't expected*)  | 
|
| 27184 | 178  | 
        | Bound 0 => instantiate' [SOME cxT] [] @{thm abs_I} (*identity: I*)
 | 
| 24827 | 179  | 
| rator$rand =>  | 
| 27184 | 180  | 
if loose_bvar1 (rator,0) then (*C or S*)  | 
| 
27179
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
181  | 
if loose_bvar1 (rand,0) then (*S*)  | 
| 
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
182  | 
let val crator = cterm_of thy (Abs(x,xT,rator))  | 
| 
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
183  | 
val crand = cterm_of thy (Abs(x,xT,rand))  | 
| 27184 | 184  | 
                     val abs_S' = cterm_instantiate [(f_S,crator),(g_S,crand)] @{thm abs_S}
 | 
185  | 
val (_,rhs) = Thm.dest_equals (cprop_of abs_S')  | 
|
| 
27179
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
186  | 
in  | 
| 
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
187  | 
Thm.transitive abs_S' (Conv.binop_conv abstract rhs)  | 
| 
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
188  | 
end  | 
| 
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
189  | 
else (*C*)  | 
| 
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
190  | 
let val crator = cterm_of thy (Abs(x,xT,rator))  | 
| 27184 | 191  | 
                     val abs_C' = cterm_instantiate [(f_C,crator),(g_C,cterm_of thy rand)] @{thm abs_C}
 | 
192  | 
val (_,rhs) = Thm.dest_equals (cprop_of abs_C')  | 
|
| 
27179
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
193  | 
in  | 
| 
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
194  | 
Thm.transitive abs_C' (Conv.fun_conv (Conv.arg_conv abstract) rhs)  | 
| 
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
195  | 
end  | 
| 27184 | 196  | 
else if loose_bvar1 (rand,0) then (*B or eta*)  | 
| 
27179
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
197  | 
if rand = Bound 0 then eta_conversion ct  | 
| 
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
198  | 
else (*B*)  | 
| 
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
199  | 
let val crand = cterm_of thy (Abs(x,xT,rand))  | 
| 
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
200  | 
val crator = cterm_of thy rator  | 
| 27184 | 201  | 
                     val abs_B' = cterm_instantiate [(f_B,crator),(g_B,crand)] @{thm abs_B}
 | 
202  | 
val (_,rhs) = Thm.dest_equals (cprop_of abs_B')  | 
|
| 
27179
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
203  | 
in  | 
| 
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
204  | 
Thm.transitive abs_B' (Conv.arg_conv abstract rhs)  | 
| 
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
205  | 
end  | 
| 
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
206  | 
else makeK()  | 
| 24827 | 207  | 
| _ => error "abstract: Bad term"  | 
208  | 
end;  | 
|
| 
20863
 
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
 
paulson 
parents: 
20789 
diff
changeset
 | 
209  | 
|
| 20419 | 210  | 
(*Traverse a theorem, declaring abstraction function definitions. String s is the suggested  | 
211  | 
prefix for the constants. Resulting theory is returned in the first theorem. *)  | 
|
| 24827 | 212  | 
fun combinators_aux ct =  | 
213  | 
if lambda_free (term_of ct) then reflexive ct  | 
|
214  | 
else  | 
|
215  | 
case term_of ct of  | 
|
216  | 
Abs _ =>  | 
|
| 
27179
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
217  | 
let val (cv,cta) = Thm.dest_abs NONE ct  | 
| 
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
218  | 
val (v,Tv) = (dest_Free o term_of) cv  | 
| 
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
219  | 
val _ = Output.debug (fn()=>" recursion: " ^ Display.string_of_cterm cta);  | 
| 
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
220  | 
val u_th = combinators_aux cta  | 
| 
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
221  | 
val _ = Output.debug (fn()=>" returned " ^ Display.string_of_thm u_th);  | 
| 
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
222  | 
val cu = Thm.rhs_of u_th  | 
| 
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
223  | 
val comb_eq = abstract (Thm.cabs cv cu)  | 
| 
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
224  | 
in Output.debug (fn()=>" abstraction result: " ^ Display.string_of_thm comb_eq);  | 
| 
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
225  | 
(transitive (abstract_rule v cv u_th) comb_eq) end  | 
| 24827 | 226  | 
| t1 $ t2 =>  | 
| 
27179
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
227  | 
let val (ct1,ct2) = Thm.dest_comb ct  | 
| 
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
228  | 
in combination (combinators_aux ct1) (combinators_aux ct2) end;  | 
| 27184 | 229  | 
|
| 24827 | 230  | 
fun combinators th =  | 
| 27184 | 231  | 
if lambda_free (prop_of th) then th  | 
| 24827 | 232  | 
else  | 
| 26928 | 233  | 
let val _ = Output.debug (fn()=>"Conversion to combinators: " ^ Display.string_of_thm th);  | 
| 
27179
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
234  | 
val th = Drule.eta_contraction_rule th  | 
| 
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
235  | 
val eqth = combinators_aux (cprop_of th)  | 
| 
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
236  | 
val _ = Output.debug (fn()=>"Conversion result: " ^ Display.string_of_thm eqth);  | 
| 
25256
 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 
paulson 
parents: 
25243 
diff
changeset
 | 
237  | 
in equal_elim eqth th end  | 
| 27184 | 238  | 
handle THM (msg,_,_) =>  | 
| 26928 | 239  | 
      (warning ("Error in the combinator translation of " ^ Display.string_of_thm th);
 | 
| 
25256
 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 
paulson 
parents: 
25243 
diff
changeset
 | 
240  | 
       warning ("  Exception message: " ^ msg);
 | 
| 
 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 
paulson 
parents: 
25243 
diff
changeset
 | 
241  | 
       TrueI);  (*A type variable of sort {} will cause make abstraction fail.*)
 | 
| 16009 | 242  | 
|
243  | 
(*cterms are used throughout for efficiency*)  | 
|
| 
18141
 
89e2e8bed08f
Skolemization by inference, but not quite finished
 
paulson 
parents: 
18009 
diff
changeset
 | 
244  | 
val cTrueprop = Thm.cterm_of HOL.thy HOLogic.Trueprop;  | 
| 16009 | 245  | 
|
246  | 
(*cterm version of mk_cTrueprop*)  | 
|
247  | 
fun c_mkTrueprop A = Thm.capply cTrueprop A;  | 
|
248  | 
||
249  | 
(*Given an abstraction over n variables, replace the bound variables by free  | 
|
250  | 
ones. Return the body, along with the list of free variables.*)  | 
|
| 
20461
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
251  | 
fun c_variant_abs_multi (ct0, vars) =  | 
| 16009 | 252  | 
let val (cv,ct) = Thm.dest_abs NONE ct0  | 
253  | 
in c_variant_abs_multi (ct, cv::vars) end  | 
|
254  | 
handle CTERM _ => (ct0, rev vars);  | 
|
255  | 
||
| 
20461
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
256  | 
(*Given the definition of a Skolem function, return a theorem to replace  | 
| 
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
257  | 
an existential formula by a use of that function.  | 
| 
18141
 
89e2e8bed08f
Skolemization by inference, but not quite finished
 
paulson 
parents: 
18009 
diff
changeset
 | 
258  | 
Example: "EX x. x : A & x ~: B ==> sko A B : A & sko A B ~: B" [.] *)  | 
| 
20461
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
259  | 
fun skolem_of_def def =  | 
| 
22902
 
ac833b4bb7ee
moved some Drule operations to Thm (see more_thm.ML);
 
wenzelm 
parents: 
22846 
diff
changeset
 | 
260  | 
let val (c,rhs) = Thm.dest_equals (cprop_of (freeze_thm def))  | 
| 16009 | 261  | 
val (ch, frees) = c_variant_abs_multi (rhs, [])  | 
| 
18141
 
89e2e8bed08f
Skolemization by inference, but not quite finished
 
paulson 
parents: 
18009 
diff
changeset
 | 
262  | 
val (chilbert,cabs) = Thm.dest_comb ch  | 
| 
26627
 
dac6d56b7c8d
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
26618 
diff
changeset
 | 
263  | 
val thy = Thm.theory_of_cterm chilbert  | 
| 
 
dac6d56b7c8d
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
26618 
diff
changeset
 | 
264  | 
val t = Thm.term_of chilbert  | 
| 
18141
 
89e2e8bed08f
Skolemization by inference, but not quite finished
 
paulson 
parents: 
18009 
diff
changeset
 | 
265  | 
      val T = case t of Const ("Hilbert_Choice.Eps", Type("fun",[_,T])) => T
 | 
| 
 
89e2e8bed08f
Skolemization by inference, but not quite finished
 
paulson 
parents: 
18009 
diff
changeset
 | 
266  | 
                      | _ => raise THM ("skolem_of_def: expected Eps", 0, [def])
 | 
| 22596 | 267  | 
val cex = Thm.cterm_of thy (HOLogic.exists_const T)  | 
| 16009 | 268  | 
val ex_tm = c_mkTrueprop (Thm.capply cex cabs)  | 
269  | 
and conc = c_mkTrueprop (Drule.beta_conv cabs (Drule.list_comb(c,frees)));  | 
|
| 
18141
 
89e2e8bed08f
Skolemization by inference, but not quite finished
 
paulson 
parents: 
18009 
diff
changeset
 | 
270  | 
fun tacf [prem] = rewrite_goals_tac [def] THEN rtac (prem RS someI_ex) 1  | 
| 23352 | 271  | 
in Goal.prove_internal [ex_tm] conc tacf  | 
| 
18141
 
89e2e8bed08f
Skolemization by inference, but not quite finished
 
paulson 
parents: 
18009 
diff
changeset
 | 
272  | 
|> forall_intr_list frees  | 
| 26653 | 273  | 
|> Thm.forall_elim_vars 0 (*Introduce Vars, but don't discharge defs.*)  | 
| 
18141
 
89e2e8bed08f
Skolemization by inference, but not quite finished
 
paulson 
parents: 
18009 
diff
changeset
 | 
274  | 
|> Thm.varifyT  | 
| 
 
89e2e8bed08f
Skolemization by inference, but not quite finished
 
paulson 
parents: 
18009 
diff
changeset
 | 
275  | 
end;  | 
| 16009 | 276  | 
|
| 
24742
 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 
paulson 
parents: 
24669 
diff
changeset
 | 
277  | 
|
| 
20863
 
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
 
paulson 
parents: 
20789 
diff
changeset
 | 
278  | 
(*Converts an Isabelle theorem (intro, elim or simp format, even higher-order) into NNF.*)  | 
| 
24937
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24854 
diff
changeset
 | 
279  | 
fun to_nnf th ctxt0 =  | 
| 
27179
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
280  | 
let val th1 = th |> transform_elim |> zero_var_indexes  | 
| 27184 | 281  | 
val ((_,[th2]),ctxt) = Variable.import_thms true [th1] ctxt0  | 
| 
24937
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24854 
diff
changeset
 | 
282  | 
val th3 = th2 |> Conv.fconv_rule ObjectLogic.atomize |> Meson.make_nnf |> strip_lambdas ~1  | 
| 
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24854 
diff
changeset
 | 
283  | 
in (th3, ctxt) end;  | 
| 16009 | 284  | 
|
| 
18141
 
89e2e8bed08f
Skolemization by inference, but not quite finished
 
paulson 
parents: 
18009 
diff
changeset
 | 
285  | 
(*Generate Skolem functions for a theorem supplied in nnf*)  | 
| 
24937
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24854 
diff
changeset
 | 
286  | 
fun assume_skolem_of_def s th =  | 
| 
22731
 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 
paulson 
parents: 
22724 
diff
changeset
 | 
287  | 
map (skolem_of_def o assume o (cterm_of (theory_of_thm th))) (assume_skofuns s th);  | 
| 
18141
 
89e2e8bed08f
Skolemization by inference, but not quite finished
 
paulson 
parents: 
18009 
diff
changeset
 | 
288  | 
|
| 24669 | 289  | 
fun assert_lambda_free ths msg =  | 
| 
20863
 
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
 
paulson 
parents: 
20789 
diff
changeset
 | 
290  | 
case filter (not o lambda_free o prop_of) ths of  | 
| 
 
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
 
paulson 
parents: 
20789 
diff
changeset
 | 
291  | 
[] => ()  | 
| 26928 | 292  | 
| ths' => error (msg ^ "\n" ^ cat_lines (map Display.string_of_thm ths'));  | 
| 
20457
 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 
paulson 
parents: 
20445 
diff
changeset
 | 
293  | 
|
| 25007 | 294  | 
|
| 27184 | 295  | 
(*** Blacklisting (duplicated in ResAtp?) ***)  | 
| 25007 | 296  | 
|
297  | 
val max_lambda_nesting = 3;  | 
|
| 27184 | 298  | 
|
| 25007 | 299  | 
fun excessive_lambdas (f$t, k) = excessive_lambdas (f,k) orelse excessive_lambdas (t,k)  | 
300  | 
| excessive_lambdas (Abs(_,_,t), k) = k=0 orelse excessive_lambdas (t,k-1)  | 
|
301  | 
| excessive_lambdas _ = false;  | 
|
302  | 
||
303  | 
fun is_formula_type T = (T = HOLogic.boolT orelse T = propT);  | 
|
304  | 
||
305  | 
(*Don't count nested lambdas at the level of formulas, as they are quantifiers*)  | 
|
306  | 
fun excessive_lambdas_fm Ts (Abs(_,T,t)) = excessive_lambdas_fm (T::Ts) t  | 
|
307  | 
| excessive_lambdas_fm Ts t =  | 
|
308  | 
if is_formula_type (fastype_of1 (Ts, t))  | 
|
309  | 
then exists (excessive_lambdas_fm Ts) (#2 (strip_comb t))  | 
|
310  | 
else excessive_lambdas (t, max_lambda_nesting);  | 
|
311  | 
||
| 
25256
 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 
paulson 
parents: 
25243 
diff
changeset
 | 
312  | 
(*The max apply_depth of any metis call in MetisExamples (on 31-10-2007) was 11.*)  | 
| 
 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 
paulson 
parents: 
25243 
diff
changeset
 | 
313  | 
val max_apply_depth = 15;  | 
| 27184 | 314  | 
|
| 
25256
 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 
paulson 
parents: 
25243 
diff
changeset
 | 
315  | 
fun apply_depth (f$t) = Int.max (apply_depth f, apply_depth t + 1)  | 
| 
 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 
paulson 
parents: 
25243 
diff
changeset
 | 
316  | 
| apply_depth (Abs(_,_,t)) = apply_depth t  | 
| 
 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 
paulson 
parents: 
25243 
diff
changeset
 | 
317  | 
| apply_depth _ = 0;  | 
| 
 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 
paulson 
parents: 
25243 
diff
changeset
 | 
318  | 
|
| 27184 | 319  | 
fun too_complex t =  | 
320  | 
apply_depth t > max_apply_depth orelse  | 
|
| 
26562
 
9d25ef112cf6
* Metis: the maximum number of clauses that can be produced from a theorem is now given by the attribute max_clauses. Theorems that exceed this number are ignored, with a warning printed.
 
paulson 
parents: 
25761 
diff
changeset
 | 
321  | 
Meson.too_many_clauses NONE t orelse  | 
| 
25256
 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 
paulson 
parents: 
25243 
diff
changeset
 | 
322  | 
excessive_lambdas_fm [] t;  | 
| 27184 | 323  | 
|
| 25243 | 324  | 
fun is_strange_thm th =  | 
325  | 
case head_of (concl_of th) of  | 
|
326  | 
Const (a,_) => (a <> "Trueprop" andalso a <> "==")  | 
|
327  | 
| _ => false;  | 
|
328  | 
||
| 27184 | 329  | 
fun bad_for_atp th =  | 
330  | 
PureThy.is_internal th  | 
|
331  | 
orelse too_complex (prop_of th)  | 
|
332  | 
orelse exists_type type_has_empty_sort (prop_of th)  | 
|
| 25761 | 333  | 
orelse is_strange_thm th;  | 
| 25243 | 334  | 
|
| 25007 | 335  | 
val multi_base_blacklist =  | 
| 
25256
 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 
paulson 
parents: 
25243 
diff
changeset
 | 
336  | 
["defs","select_defs","update_defs","induct","inducts","split","splits","split_asm",  | 
| 
 
fe467fdf129a
Catch exceptions arising during the abstraction operation.
 
paulson 
parents: 
25243 
diff
changeset
 | 
337  | 
"cases","ext_cases"]; (*FIXME: put other record thms here, or use the "Internal" marker*)  | 
| 25007 | 338  | 
|
| 21071 | 339  | 
(*Keep the full complexity of the original name*)  | 
| 
21858
 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 
wenzelm 
parents: 
21646 
diff
changeset
 | 
340  | 
fun flatten_name s = space_implode "_X" (NameSpace.explode s);  | 
| 21071 | 341  | 
|
| 
22731
 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 
paulson 
parents: 
22724 
diff
changeset
 | 
342  | 
fun fake_name th =  | 
| 24669 | 343  | 
if PureThy.has_name_hint th then flatten_name (PureThy.get_name_hint th)  | 
| 
22731
 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 
paulson 
parents: 
22724 
diff
changeset
 | 
344  | 
else gensym "unknown_thm_";  | 
| 
 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 
paulson 
parents: 
22724 
diff
changeset
 | 
345  | 
|
| 
24742
 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 
paulson 
parents: 
24669 
diff
changeset
 | 
346  | 
fun name_or_string th =  | 
| 
 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 
paulson 
parents: 
24669 
diff
changeset
 | 
347  | 
if PureThy.has_name_hint th then PureThy.get_name_hint th  | 
| 26928 | 348  | 
else Display.string_of_thm th;  | 
| 
24742
 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 
paulson 
parents: 
24669 
diff
changeset
 | 
349  | 
|
| 27184 | 350  | 
(*Skolemize a named theorem, with Skolem functions as additional premises.*)  | 
351  | 
fun skolem_thm (s, th) =  | 
|
352  | 
if member (op =) multi_base_blacklist (Sign.base_name s) orelse bad_for_atp th then []  | 
|
353  | 
else  | 
|
354  | 
let  | 
|
355  | 
val ctxt0 = Variable.thm_context th  | 
|
356  | 
val (nnfth, ctxt1) = to_nnf th ctxt0  | 
|
357  | 
val (cnfs, ctxt2) = Meson.make_cnf (assume_skolem_of_def s nnfth) nnfth ctxt1  | 
|
358  | 
in cnfs |> map combinators |> Variable.export ctxt2 ctxt0 |> Meson.finish_cnf end  | 
|
359  | 
handle THM _ => [];  | 
|
360  | 
||
| 
18510
 
0a6c24f549c3
the "skolem" attribute and better initialization of the clause database
 
paulson 
parents: 
18404 
diff
changeset
 | 
361  | 
(*Declare Skolem functions for a theorem, supplied in nnf and with its name.  | 
| 
 
0a6c24f549c3
the "skolem" attribute and better initialization of the clause database
 
paulson 
parents: 
18404 
diff
changeset
 | 
362  | 
It returns a modified theory, unless skolemization fails.*)  | 
| 27184 | 363  | 
fun skolem (name, th0) thy =  | 
| 
27179
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
364  | 
let  | 
| 
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
365  | 
val th = Thm.transfer thy th0  | 
| 
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
366  | 
val ctxt0 = Variable.thm_context th  | 
| 
24937
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24854 
diff
changeset
 | 
367  | 
in  | 
| 27184 | 368  | 
try (to_nnf th) ctxt0 |> Option.map (fn (nnfth, ctxt1) =>  | 
369  | 
let  | 
|
370  | 
val s = flatten_name name  | 
|
371  | 
val (defs, thy') = declare_skofuns s nnfth thy  | 
|
372  | 
val (cnfs, ctxt2) = Meson.make_cnf (map skolem_of_def defs) nnfth ctxt1  | 
|
373  | 
val cnfs' = cnfs |> map combinators |> Variable.export ctxt2 ctxt0  | 
|
374  | 
|> Meson.finish_cnf |> map Thm.close_derivation  | 
|
| 
27194
 
d4036ec60d46
skolem_fact/thm: uniform numbering, even for singleton list;
 
wenzelm 
parents: 
27187 
diff
changeset
 | 
375  | 
in (cnfs', thy') end)  | 
| 
24937
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24854 
diff
changeset
 | 
376  | 
end;  | 
| 16009 | 377  | 
|
| 
24742
 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 
paulson 
parents: 
24669 
diff
changeset
 | 
378  | 
(*The cache prevents repeated clausification of a theorem, and also repeated declaration of  | 
| 
 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 
paulson 
parents: 
24669 
diff
changeset
 | 
379  | 
Skolem functions.*)  | 
| 22516 | 380  | 
structure ThmCache = TheoryDataFun  | 
| 22846 | 381  | 
(  | 
| 27184 | 382  | 
type T = thm list Thmtab.table * unit Symtab.table  | 
383  | 
val empty = (Thmtab.empty, Symtab.empty)  | 
|
| 26618 | 384  | 
val copy = I;  | 
385  | 
val extend = I;  | 
|
| 27184 | 386  | 
fun merge _ ((cache1, seen1), (cache2, seen2)) : T =  | 
387  | 
(Thmtab.merge (K true) (cache1, cache2), Symtab.merge (K true) (seen1, seen2));  | 
|
| 22846 | 388  | 
);  | 
| 22516 | 389  | 
|
| 27184 | 390  | 
val lookup_cache = Thmtab.lookup o #1 o ThmCache.get;  | 
391  | 
val already_seen = Symtab.defined o #2 o ThmCache.get;  | 
|
| 
20461
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
392  | 
|
| 27184 | 393  | 
val update_cache = ThmCache.map o apfst o Thmtab.update;  | 
394  | 
fun mark_seen name = ThmCache.map (apsnd (Symtab.update (name, ())));  | 
|
| 25007 | 395  | 
|
| 
20461
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
396  | 
(*Exported function to convert Isabelle theorems into axiom clauses*)  | 
| 
27179
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
397  | 
fun cnf_axiom thy th0 =  | 
| 27184 | 398  | 
let val th = Thm.transfer thy th0 in  | 
399  | 
case lookup_cache thy th of  | 
|
400  | 
NONE => map Thm.close_derivation (skolem_thm (fake_name th, th))  | 
|
401  | 
| SOME cls => cls  | 
|
| 22516 | 402  | 
end;  | 
| 15347 | 403  | 
|
| 
18141
 
89e2e8bed08f
Skolemization by inference, but not quite finished
 
paulson 
parents: 
18009 
diff
changeset
 | 
404  | 
|
| 15872 | 405  | 
(**** Extract and Clausify theorems from a theory's claset and simpset ****)  | 
| 15347 | 406  | 
|
| 27184 | 407  | 
fun pairname th = (PureThy.get_name_hint th, th);  | 
408  | 
||
| 
17484
 
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
 
paulson 
parents: 
17412 
diff
changeset
 | 
409  | 
fun rules_of_claset cs =  | 
| 
 
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
 
paulson 
parents: 
17412 
diff
changeset
 | 
410  | 
  let val {safeIs,safeEs,hazIs,hazEs,...} = rep_cs cs
 | 
| 19175 | 411  | 
val intros = safeIs @ hazIs  | 
| 18532 | 412  | 
val elims = map Classical.classical_rule (safeEs @ hazEs)  | 
| 
17404
 
d16c3a62c396
the experimental tagging system, and the usual tidying
 
paulson 
parents: 
17279 
diff
changeset
 | 
413  | 
in  | 
| 22130 | 414  | 
Output.debug (fn () => "rules_of_claset intros: " ^ Int.toString(length intros) ^  | 
| 
17484
 
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
 
paulson 
parents: 
17412 
diff
changeset
 | 
415  | 
" elims: " ^ Int.toString(length elims));  | 
| 
20017
 
a2070352371c
made the conversion of elimination rules more robust
 
paulson 
parents: 
19894 
diff
changeset
 | 
416  | 
map pairname (intros @ elims)  | 
| 
17404
 
d16c3a62c396
the experimental tagging system, and the usual tidying
 
paulson 
parents: 
17279 
diff
changeset
 | 
417  | 
end;  | 
| 15347 | 418  | 
|
| 
17484
 
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
 
paulson 
parents: 
17412 
diff
changeset
 | 
419  | 
fun rules_of_simpset ss =  | 
| 
 
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
 
paulson 
parents: 
17412 
diff
changeset
 | 
420  | 
  let val ({rules,...}, _) = rep_ss ss
 | 
| 
 
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
 
paulson 
parents: 
17412 
diff
changeset
 | 
421  | 
val simps = Net.entries rules  | 
| 
20461
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
422  | 
in  | 
| 22130 | 423  | 
Output.debug (fn () => "rules_of_simpset: " ^ Int.toString(length simps));  | 
424  | 
map (fn r => (#name r, #thm r)) simps  | 
|
| 
17484
 
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
 
paulson 
parents: 
17412 
diff
changeset
 | 
425  | 
end;  | 
| 
 
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
 
paulson 
parents: 
17412 
diff
changeset
 | 
426  | 
|
| 21505 | 427  | 
fun claset_rules_of ctxt = rules_of_claset (local_claset_of ctxt);  | 
428  | 
fun simpset_rules_of ctxt = rules_of_simpset (local_simpset_of ctxt);  | 
|
| 
19196
 
62ee8c10d796
Added functions to retrieve local and global atpset rules.
 
mengj 
parents: 
19175 
diff
changeset
 | 
429  | 
|
| 24042 | 430  | 
fun atpset_rules_of ctxt = map pairname (ResAtpset.get ctxt);  | 
| 20774 | 431  | 
|
| 15347 | 432  | 
|
| 22471 | 433  | 
(**** Translate a set of theorems into CNF ****)  | 
| 15347 | 434  | 
|
| 19894 | 435  | 
fun pair_name_cls k (n, []) = []  | 
436  | 
| pair_name_cls k (n, cls::clss) = (cls, (n,k)) :: pair_name_cls (k+1) (n, clss)  | 
|
| 
20461
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
437  | 
|
| 
27179
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
438  | 
fun cnf_rules_pairs_aux _ pairs [] = pairs  | 
| 
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
439  | 
| cnf_rules_pairs_aux thy pairs ((name,th)::ths) =  | 
| 
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
440  | 
let val pairs' = (pair_name_cls 0 (name, cnf_axiom thy th)) @ pairs  | 
| 
20461
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
441  | 
handle THM _ => pairs | ResClause.CLAUSE _ => pairs  | 
| 
27179
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
442  | 
in cnf_rules_pairs_aux thy pairs' ths end;  | 
| 
20461
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
443  | 
|
| 
21290
 
33b6bb5d6ab8
Improvement to classrel clauses: now outputs the minimum needed.
 
paulson 
parents: 
21254 
diff
changeset
 | 
444  | 
(*The combination of rev and tail recursion preserves the original order*)  | 
| 
27179
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
445  | 
fun cnf_rules_pairs thy l = cnf_rules_pairs_aux thy [] (rev l);  | 
| 19353 | 446  | 
|
| 
19196
 
62ee8c10d796
Added functions to retrieve local and global atpset rules.
 
mengj 
parents: 
19175 
diff
changeset
 | 
447  | 
|
| 27184 | 448  | 
(**** Convert all facts of the theory into clauses (ResClause.clause, or ResHolClause.clause) ****)  | 
| 15347 | 449  | 
|
| 27184 | 450  | 
(*Populate the clause cache using the supplied theorem. Return the clausal form  | 
451  | 
and modified theory.*)  | 
|
| 
27194
 
d4036ec60d46
skolem_fact/thm: uniform numbering, even for singleton list;
 
wenzelm 
parents: 
27187 
diff
changeset
 | 
452  | 
fun skolem_cache_thm name (i, th) thy =  | 
| 27184 | 453  | 
if bad_for_atp th then thy  | 
454  | 
else  | 
|
455  | 
(case lookup_cache thy th of  | 
|
456  | 
SOME _ => thy  | 
|
457  | 
| NONE =>  | 
|
| 
27194
 
d4036ec60d46
skolem_fact/thm: uniform numbering, even for singleton list;
 
wenzelm 
parents: 
27187 
diff
changeset
 | 
458  | 
(case skolem (name ^ "_" ^ string_of_int (i + 1), th) thy of  | 
| 27184 | 459  | 
NONE => thy  | 
460  | 
| SOME (cls, thy') => update_cache (th, cls) thy'));  | 
|
| 
24742
 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 
paulson 
parents: 
24669 
diff
changeset
 | 
461  | 
|
| 27184 | 462  | 
fun skolem_cache_fact (name, ths) (changed, thy) =  | 
463  | 
if (Sign.base_name name) mem_string multi_base_blacklist orelse already_seen thy name  | 
|
464  | 
then (changed, thy)  | 
|
| 
27194
 
d4036ec60d46
skolem_fact/thm: uniform numbering, even for singleton list;
 
wenzelm 
parents: 
27187 
diff
changeset
 | 
465  | 
else (true, thy |> mark_seen name |> fold_index (skolem_cache_thm name) ths);  | 
| 
24742
 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 
paulson 
parents: 
24669 
diff
changeset
 | 
466  | 
|
| 27184 | 467  | 
fun saturate_skolem_cache thy =  | 
468  | 
(case Facts.fold_static skolem_cache_fact (PureThy.facts_of thy) (false, thy) of  | 
|
469  | 
(false, _) => NONE  | 
|
470  | 
| (true, thy') => SOME thy');  | 
|
471  | 
||
| 24854 | 472  | 
|
| 27184 | 473  | 
val suppress_endtheory = ref false;  | 
474  | 
||
475  | 
fun clause_cache_endtheory thy =  | 
|
476  | 
if ! suppress_endtheory then NONE  | 
|
477  | 
else saturate_skolem_cache thy;  | 
|
478  | 
||
| 
20457
 
85414caac94a
refinements to conversion into clause form, esp for the HO case
 
paulson 
parents: 
20445 
diff
changeset
 | 
479  | 
|
| 22516 | 480  | 
(*The cache can be kept smaller by inspecting the prop of each thm. Can ignore all that are  | 
481  | 
lambda_free, but then the individual theory caches become much bigger.*)  | 
|
| 21071 | 482  | 
|
| 
27179
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
483  | 
|
| 16563 | 484  | 
(*** meson proof methods ***)  | 
485  | 
||
| 
22731
 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 
paulson 
parents: 
22724 
diff
changeset
 | 
486  | 
(*Expand all new*definitions of abstraction or Skolem functions in a proof state.*)  | 
| 24827 | 487  | 
fun is_absko (Const ("==", _) $ Free (a,_) $ u) = String.isPrefix "sko_" a
 | 
| 
22731
 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 
paulson 
parents: 
22724 
diff
changeset
 | 
488  | 
| is_absko _ = false;  | 
| 
 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 
paulson 
parents: 
22724 
diff
changeset
 | 
489  | 
|
| 
 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 
paulson 
parents: 
22724 
diff
changeset
 | 
490  | 
fun is_okdef xs (Const ("==", _) $ t $ u) =   (*Definition of Free, not in certain terms*)
 | 
| 
 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 
paulson 
parents: 
22724 
diff
changeset
 | 
491  | 
is_Free t andalso not (member (op aconv) xs t)  | 
| 
 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 
paulson 
parents: 
22724 
diff
changeset
 | 
492  | 
| is_okdef _ _ = false  | 
| 
22724
 
3002683a6e0f
Fixes for proof reconstruction, especially involving abstractions and definitions
 
paulson 
parents: 
22691 
diff
changeset
 | 
493  | 
|
| 24215 | 494  | 
(*This function tries to cope with open locales, which introduce hypotheses of the form  | 
495  | 
Free == t, conjecture clauses, which introduce various hypotheses, and also definitions  | 
|
| 24827 | 496  | 
of sko_ functions. *)  | 
| 
22731
 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 
paulson 
parents: 
22724 
diff
changeset
 | 
497  | 
fun expand_defs_tac st0 st =  | 
| 
 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 
paulson 
parents: 
22724 
diff
changeset
 | 
498  | 
let val hyps0 = #hyps (rep_thm st0)  | 
| 
 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 
paulson 
parents: 
22724 
diff
changeset
 | 
499  | 
val hyps = #hyps (crep_thm st)  | 
| 
 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 
paulson 
parents: 
22724 
diff
changeset
 | 
500  | 
val newhyps = filter_out (member (op aconv) hyps0 o Thm.term_of) hyps  | 
| 
 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 
paulson 
parents: 
22724 
diff
changeset
 | 
501  | 
val defs = filter (is_absko o Thm.term_of) newhyps  | 
| 24669 | 502  | 
val remaining_hyps = filter_out (member (op aconv) (map Thm.term_of defs))  | 
| 
22731
 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 
paulson 
parents: 
22724 
diff
changeset
 | 
503  | 
(map Thm.term_of hyps)  | 
| 
 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 
paulson 
parents: 
22724 
diff
changeset
 | 
504  | 
val fixed = term_frees (concl_of st) @  | 
| 
 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 
paulson 
parents: 
22724 
diff
changeset
 | 
505  | 
foldl (gen_union (op aconv)) [] (map term_frees remaining_hyps)  | 
| 26928 | 506  | 
in Output.debug (fn _ => "expand_defs_tac: " ^ Display.string_of_thm st);  | 
507  | 
Output.debug (fn _ => " st0: " ^ Display.string_of_thm st0);  | 
|
508  | 
Output.debug (fn _ => " defs: " ^ commas (map Display.string_of_cterm defs));  | 
|
| 
22731
 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 
paulson 
parents: 
22724 
diff
changeset
 | 
509  | 
Seq.of_list [LocalDefs.expand (filter (is_okdef fixed o Thm.term_of) defs) st]  | 
| 
 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 
paulson 
parents: 
22724 
diff
changeset
 | 
510  | 
end;  | 
| 
22724
 
3002683a6e0f
Fixes for proof reconstruction, especially involving abstractions and definitions
 
paulson 
parents: 
22691 
diff
changeset
 | 
511  | 
|
| 
22731
 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 
paulson 
parents: 
22724 
diff
changeset
 | 
512  | 
|
| 
 
abfdccaed085
trying to  make single-step proofs work better, especially if they contain
 
paulson 
parents: 
22724 
diff
changeset
 | 
513  | 
fun meson_general_tac ths i st0 =  | 
| 
27179
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
514  | 
let  | 
| 
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
515  | 
val thy = Thm.theory_of_thm st0  | 
| 
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
516  | 
val _ = Output.debug (fn () => "Meson called: " ^ cat_lines (map Display.string_of_thm ths))  | 
| 
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
517  | 
in (Meson.meson_claset_tac (maps (cnf_axiom thy) ths) HOL_cs i THEN expand_defs_tac st0) st0 end;  | 
| 
22724
 
3002683a6e0f
Fixes for proof reconstruction, especially involving abstractions and definitions
 
paulson 
parents: 
22691 
diff
changeset
 | 
518  | 
|
| 21588 | 519  | 
val meson_method_setup = Method.add_methods  | 
520  | 
  [("meson", Method.thms_args (fn ths =>
 | 
|
| 
22724
 
3002683a6e0f
Fixes for proof reconstruction, especially involving abstractions and definitions
 
paulson 
parents: 
22691 
diff
changeset
 | 
521  | 
Method.SIMPLE_METHOD' (CHANGED_PROP o meson_general_tac ths)),  | 
| 21588 | 522  | 
"MESON resolution proof procedure")];  | 
| 15347 | 523  | 
|
| 
27179
 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 
wenzelm 
parents: 
27174 
diff
changeset
 | 
524  | 
|
| 
21999
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21900 
diff
changeset
 | 
525  | 
(*** Converting a subgoal into negated conjecture clauses. ***)  | 
| 
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21900 
diff
changeset
 | 
526  | 
|
| 24300 | 527  | 
val neg_skolemize_tac = EVERY' [rtac ccontr, ObjectLogic.atomize_prems_tac, Meson.skolemize_tac];  | 
| 22471 | 528  | 
|
| 
24937
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24854 
diff
changeset
 | 
529  | 
fun neg_clausify sts =  | 
| 
 
340523598914
context-based treatment of generalization; also handling TFrees in axiom clauses
 
paulson 
parents: 
24854 
diff
changeset
 | 
530  | 
sts |> Meson.make_clauses |> map combinators |> Meson.finish_cnf;  | 
| 
21999
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21900 
diff
changeset
 | 
531  | 
|
| 
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21900 
diff
changeset
 | 
532  | 
fun neg_conjecture_clauses st0 n =  | 
| 
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21900 
diff
changeset
 | 
533  | 
let val st = Seq.hd (neg_skolemize_tac n st0)  | 
| 
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21900 
diff
changeset
 | 
534  | 
val (params,_,_) = strip_context (Logic.nth_prem (n, Thm.prop_of st))  | 
| 27187 | 535  | 
in (neg_clausify (the (metahyps_thms n st)), params) end  | 
536  | 
handle Option => error "unable to Skolemize subgoal";  | 
|
| 
21999
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21900 
diff
changeset
 | 
537  | 
|
| 24669 | 538  | 
(*Conversion of a subgoal to conjecture clauses. Each clause has  | 
| 
21999
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21900 
diff
changeset
 | 
539  | 
leading !!-bound universal variables, to express generality. *)  | 
| 24669 | 540  | 
val neg_clausify_tac =  | 
541  | 
neg_skolemize_tac THEN'  | 
|
| 
21999
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21900 
diff
changeset
 | 
542  | 
SUBGOAL  | 
| 
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21900 
diff
changeset
 | 
543  | 
(fn (prop,_) =>  | 
| 
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21900 
diff
changeset
 | 
544  | 
let val ts = Logic.strip_assums_hyp prop  | 
| 24669 | 545  | 
in EVERY1  | 
546  | 
[METAHYPS  | 
|
547  | 
(fn hyps =>  | 
|
| 
21999
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21900 
diff
changeset
 | 
548  | 
(Method.insert_tac  | 
| 
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21900 
diff
changeset
 | 
549  | 
(map forall_intr_vars (neg_clausify hyps)) 1)),  | 
| 24669 | 550  | 
REPEAT_DETERM_N (length ts) o (etac thin_rl)]  | 
| 
21999
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21900 
diff
changeset
 | 
551  | 
end);  | 
| 
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21900 
diff
changeset
 | 
552  | 
|
| 
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21900 
diff
changeset
 | 
553  | 
val setup_methods = Method.add_methods  | 
| 24669 | 554  | 
  [("neg_clausify", Method.no_args (Method.SIMPLE_METHOD' neg_clausify_tac),
 | 
| 
21999
 
0cf192e489e2
improvements to proof reconstruction. Some files loaded in a different order
 
paulson 
parents: 
21900 
diff
changeset
 | 
555  | 
"conversion of goal to conjecture clauses")];  | 
| 24669 | 556  | 
|
| 27184 | 557  | 
|
558  | 
(** Attribute for converting a theorem into clauses **)  | 
|
559  | 
||
560  | 
val clausify = Attrib.syntax (Scan.lift Args.nat  | 
|
561  | 
>> (fn i => Thm.rule_attribute (fn context => fn th =>  | 
|
562  | 
Meson.make_meta_clause (nth (cnf_axiom (Context.theory_of context) th) i))));  | 
|
563  | 
||
564  | 
val setup_attrs = Attrib.add_attributes  | 
|
565  | 
  [("clausify", clausify, "conversion of theorem to clauses")];
 | 
|
566  | 
||
567  | 
||
568  | 
||
569  | 
(** setup **)  | 
|
570  | 
||
571  | 
val setup =  | 
|
572  | 
meson_method_setup #>  | 
|
573  | 
setup_methods #>  | 
|
574  | 
setup_attrs #>  | 
|
575  | 
perhaps saturate_skolem_cache #>  | 
|
576  | 
Theory.at_end clause_cache_endtheory;  | 
|
| 
18510
 
0a6c24f549c3
the "skolem" attribute and better initialization of the clause database
 
paulson 
parents: 
18404 
diff
changeset
 | 
577  | 
|
| 
20461
 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 
wenzelm 
parents: 
20457 
diff
changeset
 | 
578  | 
end;  | 
| 27184 | 579  |