| author | wenzelm | 
| Mon, 25 Feb 2013 17:47:32 +0100 | |
| changeset 51276 | 05522141d244 | 
| parent 49062 | 7e31dfd99ce7 | 
| child 51316 | dfe469293eb4 | 
| permissions | -rw-r--r-- | 
| 
22362
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
1  | 
(* Title: Pure/more_thm.ML  | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
2  | 
Author: Makarius  | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
3  | 
|
| 22907 | 4  | 
Further operations on type ctyp/cterm/thm, outside the inference kernel.  | 
| 
22362
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
5  | 
*)  | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
6  | 
|
| 23169 | 7  | 
infix aconvc;  | 
8  | 
||
| 32842 | 9  | 
signature BASIC_THM =  | 
10  | 
sig  | 
|
11  | 
include BASIC_THM  | 
|
12  | 
structure Ctermtab: TABLE  | 
|
13  | 
structure Thmtab: TABLE  | 
|
14  | 
val aconvc: cterm * cterm -> bool  | 
|
| 
45375
 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 
wenzelm 
parents: 
43780 
diff
changeset
 | 
15  | 
type attribute = Context.generic * thm -> Context.generic option * thm option  | 
| 32842 | 16  | 
end;  | 
17  | 
||
| 
22362
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
18  | 
signature THM =  | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
19  | 
sig  | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
20  | 
include THM  | 
| 32842 | 21  | 
structure Ctermtab: TABLE  | 
22  | 
structure Thmtab: TABLE  | 
|
| 24948 | 23  | 
val aconvc: cterm * cterm -> bool  | 
| 23491 | 24  | 
val add_cterm_frees: cterm -> cterm list -> cterm list  | 
| 32198 | 25  | 
val all_name: string * cterm -> cterm -> cterm  | 
26  | 
val all: cterm -> cterm -> cterm  | 
|
| 22907 | 27  | 
val mk_binop: cterm -> cterm -> cterm -> cterm  | 
28  | 
val dest_binop: cterm -> cterm * cterm  | 
|
29  | 
val dest_implies: cterm -> cterm * cterm  | 
|
30  | 
val dest_equals: cterm -> cterm * cterm  | 
|
31  | 
val dest_equals_lhs: cterm -> cterm  | 
|
32  | 
val dest_equals_rhs: cterm -> cterm  | 
|
33  | 
val lhs_of: thm -> cterm  | 
|
34  | 
val rhs_of: thm -> cterm  | 
|
| 
22362
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
35  | 
val thm_ord: thm * thm -> order  | 
| 32842 | 36  | 
val cterm_cache: (cterm -> 'a) -> cterm -> 'a  | 
37  | 
val thm_cache: (thm -> 'a) -> thm -> 'a  | 
|
| 23599 | 38  | 
val is_reflexive: thm -> bool  | 
| 
22362
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
39  | 
val eq_thm: thm * thm -> bool  | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
40  | 
val eq_thm_thy: thm * thm -> bool  | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
41  | 
val eq_thm_prop: thm * thm -> bool  | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
42  | 
val equiv_thm: thm * thm -> bool  | 
| 31944 | 43  | 
val class_triv: theory -> class -> thm  | 
44  | 
val of_sort: ctyp * sort -> thm list  | 
|
| 
28621
 
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
 
wenzelm 
parents: 
28116 
diff
changeset
 | 
45  | 
val check_shyps: sort list -> thm -> thm  | 
| 
24048
 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 
wenzelm 
parents: 
23599 
diff
changeset
 | 
46  | 
val is_dummy: thm -> bool  | 
| 
22695
 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 
wenzelm 
parents: 
22682 
diff
changeset
 | 
47  | 
val plain_prop_of: thm -> term  | 
| 
24048
 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 
wenzelm 
parents: 
23599 
diff
changeset
 | 
48  | 
val add_thm: thm -> thm list -> thm list  | 
| 
 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 
wenzelm 
parents: 
23599 
diff
changeset
 | 
49  | 
val del_thm: thm -> thm list -> thm list  | 
| 
 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 
wenzelm 
parents: 
23599 
diff
changeset
 | 
50  | 
val merge_thms: thm list * thm list -> thm list  | 
| 33453 | 51  | 
val full_rules: thm Item_Net.T  | 
| 30560 | 52  | 
val intro_rules: thm Item_Net.T  | 
53  | 
val elim_rules: thm Item_Net.T  | 
|
| 
27866
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
54  | 
val elim_implies: thm -> thm -> thm  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
55  | 
val forall_elim_var: int -> thm -> thm  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
56  | 
val forall_elim_vars: int -> thm -> thm  | 
| 32279 | 57  | 
val certify_inst: theory ->  | 
58  | 
((indexname * sort) * typ) list * ((indexname * typ) * term) list ->  | 
|
59  | 
(ctyp * ctyp) list * (cterm * cterm) list  | 
|
60  | 
val certify_instantiate:  | 
|
61  | 
((indexname * sort) * typ) list * ((indexname * typ) * term) list -> thm -> thm  | 
|
| 
35985
 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 
wenzelm 
parents: 
35857 
diff
changeset
 | 
62  | 
val forall_intr_frees: thm -> thm  | 
| 
35845
 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 
wenzelm 
parents: 
35715 
diff
changeset
 | 
63  | 
val unvarify_global: thm -> thm  | 
| 
27866
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
64  | 
val close_derivation: thm -> thm  | 
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
40238 
diff
changeset
 | 
65  | 
val add_axiom: Proof.context -> binding * term -> theory -> (string * thm) * theory  | 
| 
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
40238 
diff
changeset
 | 
66  | 
val add_axiom_global: binding * term -> theory -> (string * thm) * theory  | 
| 
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
40238 
diff
changeset
 | 
67  | 
val add_def: Proof.context -> bool -> bool -> binding * term -> theory -> (string * thm) * theory  | 
| 
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
40238 
diff
changeset
 | 
68  | 
val add_def_global: bool -> bool -> binding * term -> theory -> (string * thm) * theory  | 
| 
45375
 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 
wenzelm 
parents: 
43780 
diff
changeset
 | 
69  | 
type attribute = Context.generic * thm -> Context.generic option * thm option  | 
| 30210 | 70  | 
type binding = binding * attribute list  | 
71  | 
val empty_binding: binding  | 
|
| 
27866
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
72  | 
val rule_attribute: (Context.generic -> thm -> thm) -> attribute  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
73  | 
val declaration_attribute: (thm -> Context.generic -> Context.generic) -> attribute  | 
| 
45375
 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 
wenzelm 
parents: 
43780 
diff
changeset
 | 
74  | 
val mixed_attribute: (Context.generic * thm -> Context.generic * thm) -> attribute  | 
| 
46775
 
6287653e63ec
canonical argument order for attribute application;
 
wenzelm 
parents: 
46497 
diff
changeset
 | 
75  | 
val apply_attribute: attribute -> thm -> Context.generic -> thm * Context.generic  | 
| 
45375
 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 
wenzelm 
parents: 
43780 
diff
changeset
 | 
76  | 
val attribute_declaration: attribute -> thm -> Context.generic -> Context.generic  | 
| 
46775
 
6287653e63ec
canonical argument order for attribute application;
 
wenzelm 
parents: 
46497 
diff
changeset
 | 
77  | 
val theory_attributes: attribute list -> thm -> theory -> thm * theory  | 
| 
 
6287653e63ec
canonical argument order for attribute application;
 
wenzelm 
parents: 
46497 
diff
changeset
 | 
78  | 
val proof_attributes: attribute list -> thm -> Proof.context -> thm * Proof.context  | 
| 
27866
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
79  | 
val no_attributes: 'a -> 'a * 'b list  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
80  | 
  val simple_fact: 'a -> ('a * 'b list) list
 | 
| 46830 | 81  | 
val tag_rule: string * string -> thm -> thm  | 
| 
27866
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
82  | 
val untag_rule: string -> thm -> thm  | 
| 46830 | 83  | 
val tag: string * string -> attribute  | 
| 
27866
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
84  | 
val untag: string -> attribute  | 
| 30342 | 85  | 
val def_name: string -> string  | 
86  | 
val def_name_optional: string -> string -> string  | 
|
| 35238 | 87  | 
val def_binding: Binding.binding -> Binding.binding  | 
| 
30433
 
ce5138c92ca7
added def_binding_optional -- robust version of def_name_optional for bindings;
 
wenzelm 
parents: 
30342 
diff
changeset
 | 
88  | 
val def_binding_optional: Binding.binding -> Binding.binding -> Binding.binding  | 
| 
27866
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
89  | 
val has_name_hint: thm -> bool  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
90  | 
val get_name_hint: thm -> string  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
91  | 
val put_name_hint: string -> thm -> thm  | 
| 
22362
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
92  | 
val theoremK: string  | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
93  | 
val lemmaK: string  | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
94  | 
val corollaryK: string  | 
| 42473 | 95  | 
val legacy_get_kind: thm -> string  | 
| 
27866
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
96  | 
val kind_rule: string -> thm -> thm  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
97  | 
val kind: string -> attribute  | 
| 
49062
 
7e31dfd99ce7
discontinued complicated/unreliable notion of recent proofs within context;
 
wenzelm 
parents: 
49058 
diff
changeset
 | 
98  | 
val register_proofs: thm list -> theory -> theory  | 
| 
49011
 
9c68e43502ce
some support for registering forked proofs within Proof.state, using its bottom context;
 
wenzelm 
parents: 
49010 
diff
changeset
 | 
99  | 
val join_theory_proofs: theory -> unit  | 
| 
22362
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
100  | 
end;  | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
101  | 
|
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
102  | 
structure Thm: THM =  | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
103  | 
struct  | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
104  | 
|
| 
22695
 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 
wenzelm 
parents: 
22682 
diff
changeset
 | 
105  | 
(** basic operations **)  | 
| 
22362
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
106  | 
|
| 23491 | 107  | 
(* collecting cterms *)  | 
108  | 
||
109  | 
val op aconvc = op aconv o pairself Thm.term_of;  | 
|
110  | 
||
111  | 
fun add_cterm_frees ct =  | 
|
112  | 
let  | 
|
113  | 
val cert = Thm.cterm_of (Thm.theory_of_cterm ct);  | 
|
114  | 
val t = Thm.term_of ct;  | 
|
115  | 
in Term.fold_aterms (fn v as Free _ => insert (op aconvc) (cert v) | _ => I) t end;  | 
|
116  | 
||
117  | 
||
| 22907 | 118  | 
(* cterm constructors and destructors *)  | 
119  | 
||
| 32198 | 120  | 
fun all_name (x, t) A =  | 
121  | 
let  | 
|
122  | 
val cert = Thm.cterm_of (Thm.theory_of_cterm t);  | 
|
123  | 
val T = #T (Thm.rep_cterm t);  | 
|
| 
46497
 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 
wenzelm 
parents: 
45382 
diff
changeset
 | 
124  | 
  in Thm.apply (cert (Const ("all", (T --> propT) --> propT))) (Thm.lambda_name (x, t) A) end;
 | 
| 32198 | 125  | 
|
126  | 
fun all t A = all_name ("", t) A;
 | 
|
127  | 
||
| 
46497
 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 
wenzelm 
parents: 
45382 
diff
changeset
 | 
128  | 
fun mk_binop c a b = Thm.apply (Thm.apply c a) b;  | 
| 22907 | 129  | 
fun dest_binop ct = (Thm.dest_arg1 ct, Thm.dest_arg ct);  | 
130  | 
||
131  | 
fun dest_implies ct =  | 
|
132  | 
(case Thm.term_of ct of  | 
|
133  | 
    Const ("==>", _) $ _ $ _ => dest_binop ct
 | 
|
134  | 
  | _ => raise TERM ("dest_implies", [Thm.term_of ct]));
 | 
|
135  | 
||
136  | 
fun dest_equals ct =  | 
|
137  | 
(case Thm.term_of ct of  | 
|
138  | 
    Const ("==", _) $ _ $ _ => dest_binop ct
 | 
|
139  | 
  | _ => raise TERM ("dest_equals", [Thm.term_of ct]));
 | 
|
140  | 
||
141  | 
fun dest_equals_lhs ct =  | 
|
142  | 
(case Thm.term_of ct of  | 
|
143  | 
    Const ("==", _) $ _ $ _ => Thm.dest_arg1 ct
 | 
|
144  | 
  | _ => raise TERM ("dest_equals_lhs", [Thm.term_of ct]));
 | 
|
145  | 
||
146  | 
fun dest_equals_rhs ct =  | 
|
147  | 
(case Thm.term_of ct of  | 
|
148  | 
    Const ("==", _) $ _ $ _ => Thm.dest_arg ct
 | 
|
149  | 
  | _ => raise TERM ("dest_equals_rhs", [Thm.term_of ct]));
 | 
|
150  | 
||
151  | 
val lhs_of = dest_equals_lhs o Thm.cprop_of;  | 
|
152  | 
val rhs_of = dest_equals_rhs o Thm.cprop_of;  | 
|
153  | 
||
154  | 
||
155  | 
(* thm order: ignores theory context! *)  | 
|
| 22682 | 156  | 
|
| 
22362
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
157  | 
fun thm_ord (th1, th2) =  | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
158  | 
let  | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
159  | 
    val {shyps = shyps1, hyps = hyps1, tpairs = tpairs1, prop = prop1, ...} = Thm.rep_thm th1;
 | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
160  | 
    val {shyps = shyps2, hyps = hyps2, tpairs = tpairs2, prop = prop2, ...} = Thm.rep_thm th2;
 | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
161  | 
in  | 
| 35408 | 162  | 
(case Term_Ord.fast_term_ord (prop1, prop2) of  | 
| 
22362
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
163  | 
EQUAL =>  | 
| 35408 | 164  | 
(case list_ord (prod_ord Term_Ord.fast_term_ord Term_Ord.fast_term_ord) (tpairs1, tpairs2) of  | 
| 
22362
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
165  | 
EQUAL =>  | 
| 35408 | 166  | 
(case list_ord Term_Ord.fast_term_ord (hyps1, hyps2) of  | 
167  | 
EQUAL => list_ord Term_Ord.sort_ord (shyps1, shyps2)  | 
|
| 
22362
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
168  | 
| ord => ord)  | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
169  | 
| ord => ord)  | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
170  | 
| ord => ord)  | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
171  | 
end;  | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
172  | 
|
| 22682 | 173  | 
|
| 32842 | 174  | 
(* tables and caches *)  | 
175  | 
||
| 35408 | 176  | 
structure Ctermtab = Table(type key = cterm val ord = Term_Ord.fast_term_ord o pairself Thm.term_of);  | 
| 32842 | 177  | 
structure Thmtab = Table(type key = thm val ord = thm_ord);  | 
178  | 
||
179  | 
fun cterm_cache f = Cache.create Ctermtab.empty Ctermtab.lookup Ctermtab.update f;  | 
|
180  | 
fun thm_cache f = Cache.create Thmtab.empty Thmtab.lookup Thmtab.update f;  | 
|
181  | 
||
182  | 
||
| 22682 | 183  | 
(* equality *)  | 
184  | 
||
| 23599 | 185  | 
fun is_reflexive th = op aconv (Logic.dest_equals (Thm.prop_of th))  | 
186  | 
handle TERM _ => false;  | 
|
187  | 
||
| 
22362
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
188  | 
fun eq_thm ths =  | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
189  | 
Context.joinable (pairself Thm.theory_of_thm ths) andalso  | 
| 26653 | 190  | 
is_equal (thm_ord ths);  | 
| 
22362
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
191  | 
|
| 26665 | 192  | 
val eq_thm_thy = Theory.eq_thy o pairself Thm.theory_of_thm;  | 
| 
22362
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
193  | 
val eq_thm_prop = op aconv o pairself Thm.full_prop_of;  | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
194  | 
|
| 22682 | 195  | 
|
196  | 
(* pattern equivalence *)  | 
|
197  | 
||
| 
22362
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
198  | 
fun equiv_thm ths =  | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
199  | 
Pattern.equiv (Theory.merge (pairself Thm.theory_of_thm ths)) (pairself Thm.full_prop_of ths);  | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
200  | 
|
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
201  | 
|
| 
31904
 
a86896359ca4
renamed Drule.sort_triv to Thm.sort_triv (cf. more_thm.ML);
 
wenzelm 
parents: 
31177 
diff
changeset
 | 
202  | 
(* type classes and sorts *)  | 
| 
 
a86896359ca4
renamed Drule.sort_triv to Thm.sort_triv (cf. more_thm.ML);
 
wenzelm 
parents: 
31177 
diff
changeset
 | 
203  | 
|
| 31944 | 204  | 
fun class_triv thy c =  | 
205  | 
Thm.of_class (Thm.ctyp_of thy (TVar ((Name.aT, 0), [c])), c);  | 
|
206  | 
||
207  | 
fun of_sort (T, S) = map (fn c => Thm.of_class (T, c)) S;  | 
|
| 
28621
 
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
 
wenzelm 
parents: 
28116 
diff
changeset
 | 
208  | 
|
| 
 
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
 
wenzelm 
parents: 
28116 
diff
changeset
 | 
209  | 
fun check_shyps sorts raw_th =  | 
| 
 
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
 
wenzelm 
parents: 
28116 
diff
changeset
 | 
210  | 
let  | 
| 
 
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
 
wenzelm 
parents: 
28116 
diff
changeset
 | 
211  | 
val th = Thm.strip_shyps raw_th;  | 
| 
 
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
 
wenzelm 
parents: 
28116 
diff
changeset
 | 
212  | 
val prt_sort = Syntax.pretty_sort_global (Thm.theory_of_thm th);  | 
| 
 
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
 
wenzelm 
parents: 
28116 
diff
changeset
 | 
213  | 
val pending = Sorts.subtract sorts (Thm.extra_shyps th);  | 
| 
 
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
 
wenzelm 
parents: 
28116 
diff
changeset
 | 
214  | 
in  | 
| 
 
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
 
wenzelm 
parents: 
28116 
diff
changeset
 | 
215  | 
if null pending then th  | 
| 
 
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
 
wenzelm 
parents: 
28116 
diff
changeset
 | 
216  | 
else error (Pretty.string_of (Pretty.block (Pretty.str "Pending sort hypotheses:" ::  | 
| 
 
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
 
wenzelm 
parents: 
28116 
diff
changeset
 | 
217  | 
Pretty.brk 1 :: Pretty.commas (map prt_sort pending))))  | 
| 
 
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
 
wenzelm 
parents: 
28116 
diff
changeset
 | 
218  | 
end;  | 
| 
 
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
 
wenzelm 
parents: 
28116 
diff
changeset
 | 
219  | 
|
| 
 
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
 
wenzelm 
parents: 
28116 
diff
changeset
 | 
220  | 
|
| 
22695
 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 
wenzelm 
parents: 
22682 
diff
changeset
 | 
221  | 
(* misc operations *)  | 
| 
 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 
wenzelm 
parents: 
22682 
diff
changeset
 | 
222  | 
|
| 
24048
 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 
wenzelm 
parents: 
23599 
diff
changeset
 | 
223  | 
fun is_dummy thm =  | 
| 
 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 
wenzelm 
parents: 
23599 
diff
changeset
 | 
224  | 
(case try Logic.dest_term (Thm.concl_of thm) of  | 
| 
 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 
wenzelm 
parents: 
23599 
diff
changeset
 | 
225  | 
NONE => false  | 
| 
 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 
wenzelm 
parents: 
23599 
diff
changeset
 | 
226  | 
| SOME t => Term.is_dummy_pattern t);  | 
| 
 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 
wenzelm 
parents: 
23599 
diff
changeset
 | 
227  | 
|
| 
22695
 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 
wenzelm 
parents: 
22682 
diff
changeset
 | 
228  | 
fun plain_prop_of raw_thm =  | 
| 
 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 
wenzelm 
parents: 
22682 
diff
changeset
 | 
229  | 
let  | 
| 
 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 
wenzelm 
parents: 
22682 
diff
changeset
 | 
230  | 
val thm = Thm.strip_shyps raw_thm;  | 
| 
 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 
wenzelm 
parents: 
22682 
diff
changeset
 | 
231  | 
    fun err msg = raise THM ("plain_prop_of: " ^ msg, 0, [thm]);
 | 
| 
 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 
wenzelm 
parents: 
22682 
diff
changeset
 | 
232  | 
    val {hyps, prop, tpairs, ...} = Thm.rep_thm thm;
 | 
| 
 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 
wenzelm 
parents: 
22682 
diff
changeset
 | 
233  | 
in  | 
| 
 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 
wenzelm 
parents: 
22682 
diff
changeset
 | 
234  | 
if not (null hyps) then  | 
| 
 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 
wenzelm 
parents: 
22682 
diff
changeset
 | 
235  | 
err "theorem may not contain hypotheses"  | 
| 
 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 
wenzelm 
parents: 
22682 
diff
changeset
 | 
236  | 
else if not (null (Thm.extra_shyps thm)) then  | 
| 
 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 
wenzelm 
parents: 
22682 
diff
changeset
 | 
237  | 
err "theorem may not contain sort hypotheses"  | 
| 
 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 
wenzelm 
parents: 
22682 
diff
changeset
 | 
238  | 
else if not (null tpairs) then  | 
| 
 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 
wenzelm 
parents: 
22682 
diff
changeset
 | 
239  | 
err "theorem may not contain flex-flex pairs"  | 
| 
 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 
wenzelm 
parents: 
22682 
diff
changeset
 | 
240  | 
else prop  | 
| 
 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 
wenzelm 
parents: 
22682 
diff
changeset
 | 
241  | 
end;  | 
| 
 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 
wenzelm 
parents: 
22682 
diff
changeset
 | 
242  | 
|
| 
 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 
wenzelm 
parents: 
22682 
diff
changeset
 | 
243  | 
|
| 30564 | 244  | 
(* collections of theorems in canonical order *)  | 
| 
24048
 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 
wenzelm 
parents: 
23599 
diff
changeset
 | 
245  | 
|
| 
 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 
wenzelm 
parents: 
23599 
diff
changeset
 | 
246  | 
val add_thm = update eq_thm_prop;  | 
| 
 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 
wenzelm 
parents: 
23599 
diff
changeset
 | 
247  | 
val del_thm = remove eq_thm_prop;  | 
| 
 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 
wenzelm 
parents: 
23599 
diff
changeset
 | 
248  | 
val merge_thms = merge eq_thm_prop;  | 
| 
 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 
wenzelm 
parents: 
23599 
diff
changeset
 | 
249  | 
|
| 33453 | 250  | 
val full_rules = Item_Net.init eq_thm_prop (single o Thm.full_prop_of);  | 
| 33373 | 251  | 
val intro_rules = Item_Net.init eq_thm_prop (single o Thm.concl_of);  | 
252  | 
val elim_rules = Item_Net.init eq_thm_prop (single o Thm.major_prem_of);  | 
|
| 30560 | 253  | 
|
254  | 
||
| 22682 | 255  | 
|
| 
24980
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
256  | 
(** basic derived rules **)  | 
| 
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
257  | 
|
| 
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
258  | 
(*Elimination of implication  | 
| 
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
259  | 
A A ==> B  | 
| 
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
260  | 
------------  | 
| 
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
261  | 
B  | 
| 
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
262  | 
*)  | 
| 
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
263  | 
fun elim_implies thA thAB = Thm.implies_elim thAB thA;  | 
| 
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
264  | 
|
| 26653 | 265  | 
|
266  | 
(* forall_elim_var(s) *)  | 
|
267  | 
||
268  | 
local  | 
|
269  | 
||
270  | 
fun forall_elim_vars_aux strip_vars i th =  | 
|
271  | 
let  | 
|
272  | 
val thy = Thm.theory_of_thm th;  | 
|
273  | 
    val {tpairs, prop, ...} = Thm.rep_thm th;
 | 
|
274  | 
val add_used = Term.fold_aterms  | 
|
275  | 
(fn Var ((x, j), _) => if i = j then insert (op =) x else I | _ => I);  | 
|
276  | 
val used = fold (fn (t, u) => add_used t o add_used u) tpairs (add_used prop []);  | 
|
277  | 
val vars = strip_vars prop;  | 
|
278  | 
val cvars = (Name.variant_list used (map #1 vars), vars)  | 
|
279  | 
|> ListPair.map (fn (x, (_, T)) => Thm.cterm_of thy (Var ((x, i), T)));  | 
|
280  | 
in fold Thm.forall_elim cvars th end;  | 
|
281  | 
||
282  | 
in  | 
|
283  | 
||
284  | 
val forall_elim_vars = forall_elim_vars_aux Term.strip_all_vars;  | 
|
285  | 
||
| 33697 | 286  | 
fun forall_elim_var i th =  | 
287  | 
forall_elim_vars_aux  | 
|
288  | 
    (fn Const ("all", _) $ Abs (a, T, _) => [(a, T)]
 | 
|
289  | 
      | _ => raise THM ("forall_elim_vars", i, [th])) i th;
 | 
|
| 26653 | 290  | 
|
291  | 
end;  | 
|
292  | 
||
293  | 
||
| 32279 | 294  | 
(* certify_instantiate *)  | 
295  | 
||
296  | 
fun certify_inst thy (instT, inst) =  | 
|
297  | 
(map (fn (v, T) => (Thm.ctyp_of thy (TVar v), Thm.ctyp_of thy T)) instT,  | 
|
298  | 
map (fn (v, t) => (Thm.cterm_of thy (Var v), Thm.cterm_of thy t)) inst);  | 
|
299  | 
||
300  | 
fun certify_instantiate insts th =  | 
|
301  | 
Thm.instantiate (certify_inst (Thm.theory_of_thm th) insts) th;  | 
|
302  | 
||
303  | 
||
| 
35985
 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 
wenzelm 
parents: 
35857 
diff
changeset
 | 
304  | 
(* forall_intr_frees: generalization over all suitable Free variables *)  | 
| 
 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 
wenzelm 
parents: 
35857 
diff
changeset
 | 
305  | 
|
| 
 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 
wenzelm 
parents: 
35857 
diff
changeset
 | 
306  | 
fun forall_intr_frees th =  | 
| 
 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 
wenzelm 
parents: 
35857 
diff
changeset
 | 
307  | 
let  | 
| 
 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 
wenzelm 
parents: 
35857 
diff
changeset
 | 
308  | 
val thy = Thm.theory_of_thm th;  | 
| 
 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 
wenzelm 
parents: 
35857 
diff
changeset
 | 
309  | 
    val {prop, hyps, tpairs, ...} = Thm.rep_thm th;
 | 
| 
 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 
wenzelm 
parents: 
35857 
diff
changeset
 | 
310  | 
val fixed = fold Term.add_frees (Thm.terms_of_tpairs tpairs @ hyps) [];  | 
| 
 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 
wenzelm 
parents: 
35857 
diff
changeset
 | 
311  | 
val frees = Term.fold_aterms (fn Free v =>  | 
| 
 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 
wenzelm 
parents: 
35857 
diff
changeset
 | 
312  | 
if member (op =) fixed v then I else insert (op =) v | _ => I) prop [];  | 
| 
 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 
wenzelm 
parents: 
35857 
diff
changeset
 | 
313  | 
in fold (Thm.forall_intr o Thm.cterm_of thy o Free) frees th end;  | 
| 
 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 
wenzelm 
parents: 
35857 
diff
changeset
 | 
314  | 
|
| 
 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 
wenzelm 
parents: 
35857 
diff
changeset
 | 
315  | 
|
| 
35845
 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 
wenzelm 
parents: 
35715 
diff
changeset
 | 
316  | 
(* unvarify_global: global schematic variables *)  | 
| 26653 | 317  | 
|
| 
35845
 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 
wenzelm 
parents: 
35715 
diff
changeset
 | 
318  | 
fun unvarify_global th =  | 
| 
24980
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
319  | 
let  | 
| 
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
320  | 
val prop = Thm.full_prop_of th;  | 
| 
35845
 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 
wenzelm 
parents: 
35715 
diff
changeset
 | 
321  | 
val _ = map Logic.unvarify_global (prop :: Thm.hyps_of th)  | 
| 
24980
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
322  | 
handle TERM (msg, _) => raise THM (msg, 0, [th]);  | 
| 
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
323  | 
|
| 32279 | 324  | 
val instT = rev (Term.add_tvars prop []) |> map (fn v as ((a, _), S) => (v, TFree (a, S)));  | 
| 
24980
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
325  | 
val inst = rev (Term.add_vars prop []) |> map (fn ((a, i), T) =>  | 
| 32279 | 326  | 
let val T' = Term_Subst.instantiateT instT T  | 
327  | 
in (((a, i), T'), Free ((a, T'))) end);  | 
|
328  | 
in certify_instantiate (instT, inst) th end;  | 
|
| 
24980
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
329  | 
|
| 26653 | 330  | 
|
331  | 
(* close_derivation *)  | 
|
332  | 
||
| 
26628
 
63306cb94313
replaced Drule.close_derivation/Goal.close_result by Thm.close_derivation (removed obsolete compression);
 
wenzelm 
parents: 
25518 
diff
changeset
 | 
333  | 
fun close_derivation thm =  | 
| 
36744
 
6e1f3d609a68
renamed Thm.get_name -> Thm.derivation_name and Thm.put_name -> Thm.name_derivation, to emphasize the true nature of these operations;
 
wenzelm 
parents: 
36106 
diff
changeset
 | 
334  | 
if Thm.derivation_name thm = "" then Thm.name_derivation "" thm  | 
| 
26628
 
63306cb94313
replaced Drule.close_derivation/Goal.close_result by Thm.close_derivation (removed obsolete compression);
 
wenzelm 
parents: 
25518 
diff
changeset
 | 
335  | 
else thm;  | 
| 
 
63306cb94313
replaced Drule.close_derivation/Goal.close_result by Thm.close_derivation (removed obsolete compression);
 
wenzelm 
parents: 
25518 
diff
changeset
 | 
336  | 
|
| 
24980
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
337  | 
|
| 
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
338  | 
|
| 
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
339  | 
(** specification primitives **)  | 
| 
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
340  | 
|
| 30342 | 341  | 
(* rules *)  | 
342  | 
||
| 
35855
 
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
 
wenzelm 
parents: 
35853 
diff
changeset
 | 
343  | 
fun stripped_sorts thy t =  | 
| 
 
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
 
wenzelm 
parents: 
35853 
diff
changeset
 | 
344  | 
let  | 
| 
 
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
 
wenzelm 
parents: 
35853 
diff
changeset
 | 
345  | 
val tfrees = rev (map TFree (Term.add_tfrees t []));  | 
| 
43329
 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 
wenzelm 
parents: 
42473 
diff
changeset
 | 
346  | 
val tfrees' = map (fn a => TFree (a, [])) (Name.invent Name.context Name.aT (length tfrees));  | 
| 
35855
 
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
 
wenzelm 
parents: 
35853 
diff
changeset
 | 
347  | 
val strip = tfrees ~~ tfrees';  | 
| 
 
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
 
wenzelm 
parents: 
35853 
diff
changeset
 | 
348  | 
val recover = map (pairself (Thm.ctyp_of thy o Logic.varifyT_global) o swap) strip;  | 
| 
 
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
 
wenzelm 
parents: 
35853 
diff
changeset
 | 
349  | 
val t' = Term.map_types (Term.map_atyps (perhaps (AList.lookup (op =) strip))) t;  | 
| 
 
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
 
wenzelm 
parents: 
35853 
diff
changeset
 | 
350  | 
in (strip, recover, t') end;  | 
| 
 
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
 
wenzelm 
parents: 
35853 
diff
changeset
 | 
351  | 
|
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
40238 
diff
changeset
 | 
352  | 
fun add_axiom ctxt (b, prop) thy =  | 
| 
24980
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
353  | 
let  | 
| 35853 | 354  | 
    val b' = if Binding.is_empty b then Binding.name ("unnamed_axiom_" ^ serial_string ()) else b;
 | 
| 
36106
 
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
 
wenzelm 
parents: 
35988 
diff
changeset
 | 
355  | 
|
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
40238 
diff
changeset
 | 
356  | 
val _ = Sign.no_vars ctxt prop;  | 
| 
35855
 
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
 
wenzelm 
parents: 
35853 
diff
changeset
 | 
357  | 
val (strip, recover, prop') = stripped_sorts thy prop;  | 
| 
 
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
 
wenzelm 
parents: 
35853 
diff
changeset
 | 
358  | 
val constraints = map (fn (TFree (_, S), T) => (T, S)) strip;  | 
| 
 
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
 
wenzelm 
parents: 
35853 
diff
changeset
 | 
359  | 
val of_sorts = maps (fn (T as TFree (_, S), _) => of_sort (Thm.ctyp_of thy T, S)) strip;  | 
| 
36106
 
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
 
wenzelm 
parents: 
35988 
diff
changeset
 | 
360  | 
|
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
40238 
diff
changeset
 | 
361  | 
val thy' = thy  | 
| 
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
40238 
diff
changeset
 | 
362  | 
|> Theory.add_axiom ctxt (b', Logic.list_implies (maps Logic.mk_of_sort constraints, prop'));  | 
| 
36106
 
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
 
wenzelm 
parents: 
35988 
diff
changeset
 | 
363  | 
val axm_name = Sign.full_name thy' b';  | 
| 
 
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
 
wenzelm 
parents: 
35988 
diff
changeset
 | 
364  | 
val axm' = Thm.axiom thy' axm_name;  | 
| 
35988
 
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
 
wenzelm 
parents: 
35985 
diff
changeset
 | 
365  | 
val thm =  | 
| 
 
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
 
wenzelm 
parents: 
35985 
diff
changeset
 | 
366  | 
Thm.instantiate (recover, []) axm'  | 
| 
 
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
 
wenzelm 
parents: 
35985 
diff
changeset
 | 
367  | 
|> unvarify_global  | 
| 
 
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
 
wenzelm 
parents: 
35985 
diff
changeset
 | 
368  | 
|> fold elim_implies of_sorts;  | 
| 
36106
 
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
 
wenzelm 
parents: 
35988 
diff
changeset
 | 
369  | 
in ((axm_name, thm), thy') end;  | 
| 
24980
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
370  | 
|
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
40238 
diff
changeset
 | 
371  | 
fun add_axiom_global arg thy = add_axiom (Syntax.init_pretty_global thy) arg thy;  | 
| 
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
40238 
diff
changeset
 | 
372  | 
|
| 
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
40238 
diff
changeset
 | 
373  | 
fun add_def ctxt unchecked overloaded (b, prop) thy =  | 
| 
24980
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
374  | 
let  | 
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
40238 
diff
changeset
 | 
375  | 
val _ = Sign.no_vars ctxt prop;  | 
| 
35988
 
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
 
wenzelm 
parents: 
35985 
diff
changeset
 | 
376  | 
val prems = map (Thm.cterm_of thy) (Logic.strip_imp_prems prop);  | 
| 
 
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
 
wenzelm 
parents: 
35985 
diff
changeset
 | 
377  | 
val (_, recover, concl') = stripped_sorts thy (Logic.strip_imp_concl prop);  | 
| 
36106
 
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
 
wenzelm 
parents: 
35988 
diff
changeset
 | 
378  | 
|
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
40238 
diff
changeset
 | 
379  | 
val thy' = Theory.add_def ctxt unchecked overloaded (b, concl') thy;  | 
| 
36106
 
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
 
wenzelm 
parents: 
35988 
diff
changeset
 | 
380  | 
val axm_name = Sign.full_name thy' b;  | 
| 
 
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
 
wenzelm 
parents: 
35988 
diff
changeset
 | 
381  | 
val axm' = Thm.axiom thy' axm_name;  | 
| 
35988
 
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
 
wenzelm 
parents: 
35985 
diff
changeset
 | 
382  | 
val thm =  | 
| 
 
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
 
wenzelm 
parents: 
35985 
diff
changeset
 | 
383  | 
Thm.instantiate (recover, []) axm'  | 
| 
 
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
 
wenzelm 
parents: 
35985 
diff
changeset
 | 
384  | 
|> unvarify_global  | 
| 
 
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
 
wenzelm 
parents: 
35985 
diff
changeset
 | 
385  | 
|> fold_rev Thm.implies_intr prems;  | 
| 
36106
 
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
 
wenzelm 
parents: 
35988 
diff
changeset
 | 
386  | 
in ((axm_name, thm), thy') end;  | 
| 
24980
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
387  | 
|
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
40238 
diff
changeset
 | 
388  | 
fun add_def_global unchecked overloaded arg thy =  | 
| 
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
40238 
diff
changeset
 | 
389  | 
add_def (Syntax.init_pretty_global thy) unchecked overloaded arg thy;  | 
| 
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
40238 
diff
changeset
 | 
390  | 
|
| 
27866
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
391  | 
|
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
392  | 
|
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
393  | 
(** attributes **)  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
394  | 
|
| 
40238
 
edcdecd55655
type attribute is derived concept outside the kernel;
 
wenzelm 
parents: 
39133 
diff
changeset
 | 
395  | 
(*attributes subsume any kind of rules or context modifiers*)  | 
| 
45375
 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 
wenzelm 
parents: 
43780 
diff
changeset
 | 
396  | 
type attribute = Context.generic * thm -> Context.generic option * thm option;  | 
| 
40238
 
edcdecd55655
type attribute is derived concept outside the kernel;
 
wenzelm 
parents: 
39133 
diff
changeset
 | 
397  | 
|
| 30210 | 398  | 
type binding = binding * attribute list;  | 
399  | 
val empty_binding: binding = (Binding.empty, []);  | 
|
400  | 
||
| 
45375
 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 
wenzelm 
parents: 
43780 
diff
changeset
 | 
401  | 
fun rule_attribute f (x, th) = (NONE, SOME (f x th));  | 
| 
 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 
wenzelm 
parents: 
43780 
diff
changeset
 | 
402  | 
fun declaration_attribute f (x, th) = (SOME (f th x), NONE);  | 
| 
 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 
wenzelm 
parents: 
43780 
diff
changeset
 | 
403  | 
fun mixed_attribute f (x, th) = let val (x', th') = f (x, th) in (SOME x', SOME th') end;  | 
| 
 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 
wenzelm 
parents: 
43780 
diff
changeset
 | 
404  | 
|
| 
46775
 
6287653e63ec
canonical argument order for attribute application;
 
wenzelm 
parents: 
46497 
diff
changeset
 | 
405  | 
fun apply_attribute (att: attribute) th x =  | 
| 
45375
 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 
wenzelm 
parents: 
43780 
diff
changeset
 | 
406  | 
let val (x', th') = att (x, th)  | 
| 
46775
 
6287653e63ec
canonical argument order for attribute application;
 
wenzelm 
parents: 
46497 
diff
changeset
 | 
407  | 
in (the_default th th', the_default x x') end;  | 
| 
45375
 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 
wenzelm 
parents: 
43780 
diff
changeset
 | 
408  | 
|
| 
46775
 
6287653e63ec
canonical argument order for attribute application;
 
wenzelm 
parents: 
46497 
diff
changeset
 | 
409  | 
fun attribute_declaration att th x = #2 (apply_attribute att th x);  | 
| 
27866
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
410  | 
|
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
411  | 
fun apply_attributes mk dest =  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
412  | 
let  | 
| 
46775
 
6287653e63ec
canonical argument order for attribute application;
 
wenzelm 
parents: 
46497 
diff
changeset
 | 
413  | 
fun app [] th x = (th, x)  | 
| 
 
6287653e63ec
canonical argument order for attribute application;
 
wenzelm 
parents: 
46497 
diff
changeset
 | 
414  | 
| app (att :: atts) th x = apply_attribute att th (mk x) ||> dest |-> app atts;  | 
| 
27866
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
415  | 
in app end;  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
416  | 
|
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
417  | 
val theory_attributes = apply_attributes Context.Theory Context.the_theory;  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
418  | 
val proof_attributes = apply_attributes Context.Proof Context.the_proof;  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
419  | 
|
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
420  | 
fun no_attributes x = (x, []);  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
421  | 
fun simple_fact x = [(x, [])];  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
422  | 
|
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
423  | 
|
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
424  | 
|
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
425  | 
(*** theorem tags ***)  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
426  | 
|
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
427  | 
(* add / delete tags *)  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
428  | 
|
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
429  | 
fun tag_rule tg = Thm.map_tags (insert (op =) tg);  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
430  | 
fun untag_rule s = Thm.map_tags (filter_out (fn (s', _) => s = s'));  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
431  | 
|
| 
45375
 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 
wenzelm 
parents: 
43780 
diff
changeset
 | 
432  | 
fun tag tg = rule_attribute (K (tag_rule tg));  | 
| 
 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 
wenzelm 
parents: 
43780 
diff
changeset
 | 
433  | 
fun untag s = rule_attribute (K (untag_rule s));  | 
| 
27866
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
434  | 
|
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
435  | 
|
| 30342 | 436  | 
(* def_name *)  | 
437  | 
||
438  | 
fun def_name c = c ^ "_def";  | 
|
439  | 
||
440  | 
fun def_name_optional c "" = def_name c  | 
|
441  | 
| def_name_optional _ name = name;  | 
|
442  | 
||
| 35238 | 443  | 
val def_binding = Binding.map_name def_name;  | 
444  | 
||
| 
30433
 
ce5138c92ca7
added def_binding_optional -- robust version of def_name_optional for bindings;
 
wenzelm 
parents: 
30342 
diff
changeset
 | 
445  | 
fun def_binding_optional b name =  | 
| 35238 | 446  | 
if Binding.is_empty name then def_binding b else name;  | 
| 
30433
 
ce5138c92ca7
added def_binding_optional -- robust version of def_name_optional for bindings;
 
wenzelm 
parents: 
30342 
diff
changeset
 | 
447  | 
|
| 30342 | 448  | 
|
| 
27866
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
449  | 
(* unofficial theorem names *)  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
450  | 
|
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
451  | 
fun the_name_hint thm = the (AList.lookup (op =) (Thm.get_tags thm) Markup.nameN);  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
452  | 
|
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
453  | 
val has_name_hint = can the_name_hint;  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
454  | 
val get_name_hint = the_default "??.unknown" o try the_name_hint;  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
455  | 
|
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
456  | 
fun put_name_hint name = untag_rule Markup.nameN #> tag_rule (Markup.nameN, name);  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
457  | 
|
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
458  | 
|
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
459  | 
(* theorem kinds *)  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
460  | 
|
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
461  | 
val theoremK = "theorem";  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
462  | 
val lemmaK = "lemma";  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
463  | 
val corollaryK = "corollary";  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
464  | 
|
| 42473 | 465  | 
fun legacy_get_kind thm = the_default "" (Properties.get (Thm.get_tags thm) Markup.kindN);  | 
| 
27866
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
466  | 
|
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
467  | 
fun kind_rule k = tag_rule (Markup.kindN, k) o untag_rule Markup.kindN;  | 
| 
45375
 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 
wenzelm 
parents: 
43780 
diff
changeset
 | 
468  | 
fun kind k = rule_attribute (K (k <> "" ? kind_rule k));  | 
| 
27866
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
469  | 
|
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
470  | 
|
| 
49011
 
9c68e43502ce
some support for registering forked proofs within Proof.state, using its bottom context;
 
wenzelm 
parents: 
49010 
diff
changeset
 | 
471  | 
(* forked proofs *)  | 
| 49010 | 472  | 
|
| 
49062
 
7e31dfd99ce7
discontinued complicated/unreliable notion of recent proofs within context;
 
wenzelm 
parents: 
49058 
diff
changeset
 | 
473  | 
structure Proofs = Theory_Data  | 
| 49010 | 474  | 
(  | 
| 
49062
 
7e31dfd99ce7
discontinued complicated/unreliable notion of recent proofs within context;
 
wenzelm 
parents: 
49058 
diff
changeset
 | 
475  | 
type T = thm list;  | 
| 
 
7e31dfd99ce7
discontinued complicated/unreliable notion of recent proofs within context;
 
wenzelm 
parents: 
49058 
diff
changeset
 | 
476  | 
val empty = [];  | 
| 49010 | 477  | 
fun extend _ = empty;  | 
478  | 
fun merge _ = empty;  | 
|
479  | 
);  | 
|
480  | 
||
| 
49062
 
7e31dfd99ce7
discontinued complicated/unreliable notion of recent proofs within context;
 
wenzelm 
parents: 
49058 
diff
changeset
 | 
481  | 
fun register_proofs more_thms = Proofs.map (fn thms => fold cons more_thms thms);  | 
| 
 
7e31dfd99ce7
discontinued complicated/unreliable notion of recent proofs within context;
 
wenzelm 
parents: 
49058 
diff
changeset
 | 
482  | 
val join_theory_proofs = Thm.join_proofs o rev o Proofs.get;  | 
| 49010 | 483  | 
|
484  | 
||
| 
22362
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
485  | 
open Thm;  | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
486  | 
|
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
487  | 
end;  | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
488  | 
|
| 32842 | 489  | 
structure Basic_Thm: BASIC_THM = Thm;  | 
490  | 
open Basic_Thm;  | 
|
| 23170 | 491  |