| author | wenzelm | 
| Tue, 02 Aug 2016 17:39:38 +0200 | |
| changeset 63580 | 7f06347a5013 | 
| parent 63352 | 4eaf35781b23 | 
| child 64556 | 851ae0e7b09c | 
| 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  | 
|
| 61268 | 12  | 
val show_consts: bool Config.T  | 
13  | 
val show_hyps: bool Config.T  | 
|
14  | 
val show_tags: bool Config.T  | 
|
| 32842 | 15  | 
structure Ctermtab: TABLE  | 
16  | 
structure Thmtab: TABLE  | 
|
17  | 
val aconvc: cterm * cterm -> bool  | 
|
| 
45375
 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 
wenzelm 
parents: 
43780 
diff
changeset
 | 
18  | 
type attribute = Context.generic * thm -> Context.generic option * thm option  | 
| 32842 | 19  | 
end;  | 
20  | 
||
| 
22362
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
21  | 
signature THM =  | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
22  | 
sig  | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
23  | 
include THM  | 
| 32842 | 24  | 
structure Ctermtab: TABLE  | 
25  | 
structure Thmtab: TABLE  | 
|
| 
60952
 
762cb38a3147
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60951 
diff
changeset
 | 
26  | 
val eq_ctyp: ctyp * ctyp -> bool  | 
| 24948 | 27  | 
val aconvc: cterm * cterm -> bool  | 
| 
60952
 
762cb38a3147
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60951 
diff
changeset
 | 
28  | 
val add_tvars: thm -> ctyp list -> ctyp list  | 
| 60818 | 29  | 
val add_frees: thm -> cterm list -> cterm list  | 
30  | 
val add_vars: thm -> cterm list -> cterm list  | 
|
| 60938 | 31  | 
val all_name: Proof.context -> string * cterm -> cterm -> cterm  | 
32  | 
val all: Proof.context -> cterm -> cterm -> cterm  | 
|
| 22907 | 33  | 
val mk_binop: cterm -> cterm -> cterm -> cterm  | 
34  | 
val dest_binop: cterm -> cterm * cterm  | 
|
35  | 
val dest_implies: cterm -> cterm * cterm  | 
|
36  | 
val dest_equals: cterm -> cterm * cterm  | 
|
37  | 
val dest_equals_lhs: cterm -> cterm  | 
|
38  | 
val dest_equals_rhs: cterm -> cterm  | 
|
39  | 
val lhs_of: thm -> cterm  | 
|
40  | 
val rhs_of: thm -> cterm  | 
|
| 
22362
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
41  | 
val thm_ord: thm * thm -> order  | 
| 32842 | 42  | 
val cterm_cache: (cterm -> 'a) -> cterm -> 'a  | 
43  | 
val thm_cache: (thm -> 'a) -> thm -> 'a  | 
|
| 23599 | 44  | 
val is_reflexive: thm -> bool  | 
| 
22362
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
45  | 
val eq_thm: thm * thm -> bool  | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
46  | 
val eq_thm_prop: thm * thm -> bool  | 
| 
52683
 
fb028440473e
more official Thm.eq_thm_strict, without demanding ML equality type;
 
wenzelm 
parents: 
51316 
diff
changeset
 | 
47  | 
val eq_thm_strict: thm * thm -> bool  | 
| 60817 | 48  | 
val equiv_thm: theory -> thm * thm -> bool  | 
| 31944 | 49  | 
val class_triv: theory -> class -> thm  | 
50  | 
val of_sort: ctyp * sort -> thm list  | 
|
| 
24048
 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 
wenzelm 
parents: 
23599 
diff
changeset
 | 
51  | 
val is_dummy: thm -> bool  | 
| 
22695
 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 
wenzelm 
parents: 
22682 
diff
changeset
 | 
52  | 
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
 | 
53  | 
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
 | 
54  | 
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
 | 
55  | 
val merge_thms: thm list * thm list -> thm list  | 
| 33453 | 56  | 
val full_rules: thm Item_Net.T  | 
| 30560 | 57  | 
val intro_rules: thm Item_Net.T  | 
58  | 
val elim_rules: thm Item_Net.T  | 
|
| 
54984
 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 
wenzelm 
parents: 
53206 
diff
changeset
 | 
59  | 
val declare_hyps: cterm -> Proof.context -> Proof.context  | 
| 
 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 
wenzelm 
parents: 
53206 
diff
changeset
 | 
60  | 
val assume_hyps: cterm -> Proof.context -> thm * Proof.context  | 
| 
54993
 
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
 
wenzelm 
parents: 
54984 
diff
changeset
 | 
61  | 
val unchecked_hyps: Proof.context -> Proof.context  | 
| 
 
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
 
wenzelm 
parents: 
54984 
diff
changeset
 | 
62  | 
val restore_hyps: Proof.context -> Proof.context -> Proof.context  | 
| 55633 | 63  | 
val undeclared_hyps: Context.generic -> thm -> term list  | 
| 
54993
 
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
 
wenzelm 
parents: 
54984 
diff
changeset
 | 
64  | 
val check_hyps: Context.generic -> thm -> thm  | 
| 61508 | 65  | 
val declare_term_sorts: term -> Proof.context -> Proof.context  | 
| 
61509
 
358dfae15d83
print thm wrt. local shyps (from full proof context);
 
wenzelm 
parents: 
61508 
diff
changeset
 | 
66  | 
val extra_shyps': Proof.context -> thm -> sort list  | 
| 61508 | 67  | 
val check_shyps: Proof.context -> thm -> thm  | 
68  | 
val weaken_sorts': Proof.context -> cterm -> cterm  | 
|
| 
27866
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
69  | 
val elim_implies: thm -> thm -> thm  | 
| 61339 | 70  | 
val forall_intr_name: string * cterm -> thm -> thm  | 
| 
27866
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
71  | 
val forall_elim_var: int -> thm -> thm  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
72  | 
val forall_elim_vars: int -> thm -> thm  | 
| 60801 | 73  | 
val instantiate': ctyp option list -> cterm option 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
 | 
74  | 
val forall_intr_frees: thm -> thm  | 
| 60825 | 75  | 
val unvarify_global: theory -> thm -> thm  | 
76  | 
val unvarify_axiom: theory -> string -> thm  | 
|
| 
27866
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
77  | 
val close_derivation: thm -> thm  | 
| 
59969
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
78  | 
val rename_params_rule: string list * int -> thm -> thm  | 
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
79  | 
val rename_boundvars: term -> term -> thm -> thm  | 
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
40238 
diff
changeset
 | 
80  | 
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
 | 
81  | 
val add_axiom_global: binding * term -> theory -> (string * thm) * theory  | 
| 
61261
 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 
wenzelm 
parents: 
61059 
diff
changeset
 | 
82  | 
val add_def: Defs.context -> bool -> bool -> binding * term -> theory -> (string * thm) * theory  | 
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
40238 
diff
changeset
 | 
83  | 
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
 | 
84  | 
type attribute = Context.generic * thm -> Context.generic option * thm option  | 
| 30210 | 85  | 
type binding = binding * attribute list  | 
| 46830 | 86  | 
val tag_rule: string * string -> thm -> thm  | 
| 
27866
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
87  | 
val untag_rule: string -> thm -> thm  | 
| 61852 | 88  | 
val is_free_dummy: thm -> bool  | 
89  | 
val tag_free_dummy: thm -> thm  | 
|
| 30342 | 90  | 
val def_name: string -> string  | 
91  | 
val def_name_optional: string -> string -> string  | 
|
| 35238 | 92  | 
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
 | 
93  | 
val def_binding_optional: Binding.binding -> Binding.binding -> Binding.binding  | 
| 62093 | 94  | 
val make_def_binding: bool -> Binding.binding -> Binding.binding  | 
| 
27866
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
95  | 
val has_name_hint: thm -> bool  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
96  | 
val get_name_hint: thm -> string  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
97  | 
val put_name_hint: string -> thm -> thm  | 
| 
22362
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
98  | 
val theoremK: string  | 
| 42473 | 99  | 
val legacy_get_kind: thm -> string  | 
| 
27866
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
100  | 
val kind_rule: string -> thm -> thm  | 
| 
61853
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
101  | 
val rule_attribute: thm list -> (Context.generic -> thm -> thm) -> attribute  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
102  | 
val declaration_attribute: (thm -> Context.generic -> Context.generic) -> attribute  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
103  | 
val mixed_attribute: (Context.generic * thm -> Context.generic * thm) -> attribute  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
104  | 
val apply_attribute: attribute -> thm -> Context.generic -> thm * Context.generic  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
105  | 
val attribute_declaration: attribute -> thm -> Context.generic -> Context.generic  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
106  | 
val theory_attributes: attribute list -> thm -> theory -> thm * theory  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
107  | 
val proof_attributes: attribute list -> thm -> Proof.context -> thm * Proof.context  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
108  | 
val no_attributes: 'a -> 'a * 'b list  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
109  | 
  val simple_fact: 'a -> ('a * 'b list) list
 | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
110  | 
val tag: string * string -> attribute  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
111  | 
val untag: string -> attribute  | 
| 
27866
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
112  | 
val kind: string -> attribute  | 
| 
49062
 
7e31dfd99ce7
discontinued complicated/unreliable notion of recent proofs within context;
 
wenzelm 
parents: 
49058 
diff
changeset
 | 
113  | 
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
 | 
114  | 
val join_theory_proofs: theory -> unit  | 
| 61268 | 115  | 
val show_consts_raw: Config.raw  | 
116  | 
val show_consts: bool Config.T  | 
|
117  | 
val show_hyps_raw: Config.raw  | 
|
118  | 
val show_hyps: bool Config.T  | 
|
119  | 
val show_tags_raw: Config.raw  | 
|
120  | 
val show_tags: bool Config.T  | 
|
121  | 
val pretty_flexpair: Proof.context -> term * term -> Pretty.T  | 
|
122  | 
  val pretty_thm_raw: Proof.context -> {quote: bool, show_hyps: bool} -> thm -> Pretty.T
 | 
|
123  | 
val pretty_thm: Proof.context -> thm -> Pretty.T  | 
|
124  | 
val pretty_thm_item: Proof.context -> thm -> Pretty.T  | 
|
125  | 
val pretty_thm_global: theory -> thm -> Pretty.T  | 
|
126  | 
val string_of_thm: Proof.context -> thm -> string  | 
|
127  | 
val string_of_thm_global: theory -> thm -> string  | 
|
| 
22362
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
128  | 
end;  | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
129  | 
|
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
130  | 
structure Thm: THM =  | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
131  | 
struct  | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
132  | 
|
| 
22695
 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 
wenzelm 
parents: 
22682 
diff
changeset
 | 
133  | 
(** basic operations **)  | 
| 
22362
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
134  | 
|
| 
60952
 
762cb38a3147
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60951 
diff
changeset
 | 
135  | 
(* collecting ctyps and cterms *)  | 
| 23491 | 136  | 
|
| 
60952
 
762cb38a3147
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60951 
diff
changeset
 | 
137  | 
val eq_ctyp = op = o apply2 Thm.typ_of;  | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58001 
diff
changeset
 | 
138  | 
val op aconvc = op aconv o apply2 Thm.term_of;  | 
| 23491 | 139  | 
|
| 
60952
 
762cb38a3147
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60951 
diff
changeset
 | 
140  | 
val add_tvars = Thm.fold_atomic_ctyps (fn a => is_TVar (Thm.typ_of a) ? insert eq_ctyp a);  | 
| 60818 | 141  | 
val add_frees = Thm.fold_atomic_cterms (fn a => is_Free (Thm.term_of a) ? insert (op aconvc) a);  | 
142  | 
val add_vars = Thm.fold_atomic_cterms (fn a => is_Var (Thm.term_of a) ? insert (op aconvc) a);  | 
|
| 23491 | 143  | 
|
144  | 
||
| 22907 | 145  | 
(* cterm constructors and destructors *)  | 
146  | 
||
| 60938 | 147  | 
fun all_name ctxt (x, t) A =  | 
| 32198 | 148  | 
let  | 
| 59586 | 149  | 
val T = Thm.typ_of_cterm t;  | 
| 60938 | 150  | 
    val all_const = Thm.cterm_of ctxt (Const ("Pure.all", (T --> propT) --> propT));
 | 
151  | 
in Thm.apply all_const (Thm.lambda_name (x, t) A) end;  | 
|
| 32198 | 152  | 
|
| 60938 | 153  | 
fun all ctxt t A = all_name ctxt ("", t) A;
 | 
| 32198 | 154  | 
|
| 
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
 | 
155  | 
fun mk_binop c a b = Thm.apply (Thm.apply c a) b;  | 
| 22907 | 156  | 
fun dest_binop ct = (Thm.dest_arg1 ct, Thm.dest_arg ct);  | 
157  | 
||
158  | 
fun dest_implies ct =  | 
|
159  | 
(case Thm.term_of ct of  | 
|
| 56245 | 160  | 
    Const ("Pure.imp", _) $ _ $ _ => dest_binop ct
 | 
| 22907 | 161  | 
  | _ => raise TERM ("dest_implies", [Thm.term_of ct]));
 | 
162  | 
||
163  | 
fun dest_equals ct =  | 
|
164  | 
(case Thm.term_of ct of  | 
|
| 56245 | 165  | 
    Const ("Pure.eq", _) $ _ $ _ => dest_binop ct
 | 
| 22907 | 166  | 
  | _ => raise TERM ("dest_equals", [Thm.term_of ct]));
 | 
167  | 
||
168  | 
fun dest_equals_lhs ct =  | 
|
169  | 
(case Thm.term_of ct of  | 
|
| 56245 | 170  | 
    Const ("Pure.eq", _) $ _ $ _ => Thm.dest_arg1 ct
 | 
| 22907 | 171  | 
  | _ => raise TERM ("dest_equals_lhs", [Thm.term_of ct]));
 | 
172  | 
||
173  | 
fun dest_equals_rhs ct =  | 
|
174  | 
(case Thm.term_of ct of  | 
|
| 56245 | 175  | 
    Const ("Pure.eq", _) $ _ $ _ => Thm.dest_arg ct
 | 
| 22907 | 176  | 
  | _ => raise TERM ("dest_equals_rhs", [Thm.term_of ct]));
 | 
177  | 
||
178  | 
val lhs_of = dest_equals_lhs o Thm.cprop_of;  | 
|
179  | 
val rhs_of = dest_equals_rhs o Thm.cprop_of;  | 
|
180  | 
||
181  | 
||
182  | 
(* thm order: ignores theory context! *)  | 
|
| 22682 | 183  | 
|
| 61039 | 184  | 
fun thm_ord ths =  | 
185  | 
(case Term_Ord.fast_term_ord (apply2 Thm.prop_of ths) of  | 
|
186  | 
EQUAL =>  | 
|
187  | 
(case  | 
|
188  | 
list_ord (prod_ord Term_Ord.fast_term_ord Term_Ord.fast_term_ord)  | 
|
189  | 
(apply2 Thm.tpairs_of ths)  | 
|
190  | 
of  | 
|
191  | 
EQUAL =>  | 
|
192  | 
(case list_ord Term_Ord.fast_term_ord (apply2 Thm.hyps_of ths) of  | 
|
193  | 
EQUAL => list_ord Term_Ord.sort_ord (apply2 Thm.shyps_of ths)  | 
|
194  | 
| ord => ord)  | 
|
195  | 
| ord => ord)  | 
|
196  | 
| ord => ord);  | 
|
| 
22362
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
197  | 
|
| 22682 | 198  | 
|
| 32842 | 199  | 
(* tables and caches *)  | 
200  | 
||
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58001 
diff
changeset
 | 
201  | 
structure Ctermtab = Table(type key = cterm val ord = Term_Ord.fast_term_ord o apply2 Thm.term_of);  | 
| 32842 | 202  | 
structure Thmtab = Table(type key = thm val ord = thm_ord);  | 
203  | 
||
204  | 
fun cterm_cache f = Cache.create Ctermtab.empty Ctermtab.lookup Ctermtab.update f;  | 
|
205  | 
fun thm_cache f = Cache.create Thmtab.empty Thmtab.lookup Thmtab.update f;  | 
|
206  | 
||
207  | 
||
| 22682 | 208  | 
(* equality *)  | 
209  | 
||
| 23599 | 210  | 
fun is_reflexive th = op aconv (Logic.dest_equals (Thm.prop_of th))  | 
211  | 
handle TERM _ => false;  | 
|
212  | 
||
| 
55547
 
384bfd19ee61
subtle change of semantics of Thm.eq_thm, e.g. relevant for merge of src/HOL/Tools/Predicate_Compile/core_data.ML (cf. HOL-IMP);
 
wenzelm 
parents: 
54996 
diff
changeset
 | 
213  | 
val eq_thm = is_equal o thm_ord;  | 
| 
22362
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
214  | 
|
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58001 
diff
changeset
 | 
215  | 
val eq_thm_prop = op aconv o apply2 Thm.full_prop_of;  | 
| 
22362
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
216  | 
|
| 
52683
 
fb028440473e
more official Thm.eq_thm_strict, without demanding ML equality type;
 
wenzelm 
parents: 
51316 
diff
changeset
 | 
217  | 
fun eq_thm_strict ths =  | 
| 
55547
 
384bfd19ee61
subtle change of semantics of Thm.eq_thm, e.g. relevant for merge of src/HOL/Tools/Predicate_Compile/core_data.ML (cf. HOL-IMP);
 
wenzelm 
parents: 
54996 
diff
changeset
 | 
218  | 
eq_thm ths andalso  | 
| 61040 | 219  | 
Context.eq_thy_id (apply2 Thm.theory_id_of_thm ths) andalso  | 
220  | 
op = (apply2 Thm.maxidx_of ths) andalso  | 
|
221  | 
op = (apply2 Thm.get_tags ths);  | 
|
| 
52683
 
fb028440473e
more official Thm.eq_thm_strict, without demanding ML equality type;
 
wenzelm 
parents: 
51316 
diff
changeset
 | 
222  | 
|
| 22682 | 223  | 
|
224  | 
(* pattern equivalence *)  | 
|
225  | 
||
| 60817 | 226  | 
fun equiv_thm thy ths =  | 
227  | 
Pattern.equiv thy (apply2 (Thm.full_prop_of o Thm.transfer thy) ths);  | 
|
| 
22362
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
228  | 
|
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
229  | 
|
| 
31904
 
a86896359ca4
renamed Drule.sort_triv to Thm.sort_triv (cf. more_thm.ML);
 
wenzelm 
parents: 
31177 
diff
changeset
 | 
230  | 
(* type classes and sorts *)  | 
| 
 
a86896359ca4
renamed Drule.sort_triv to Thm.sort_triv (cf. more_thm.ML);
 
wenzelm 
parents: 
31177 
diff
changeset
 | 
231  | 
|
| 31944 | 232  | 
fun class_triv thy c =  | 
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59616 
diff
changeset
 | 
233  | 
Thm.of_class (Thm.global_ctyp_of thy (TVar ((Name.aT, 0), [c])), c);  | 
| 31944 | 234  | 
|
235  | 
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
 | 
236  | 
|
| 
 
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
 
wenzelm 
parents: 
28116 
diff
changeset
 | 
237  | 
|
| 
22695
 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 
wenzelm 
parents: 
22682 
diff
changeset
 | 
238  | 
(* misc operations *)  | 
| 
 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 
wenzelm 
parents: 
22682 
diff
changeset
 | 
239  | 
|
| 
24048
 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 
wenzelm 
parents: 
23599 
diff
changeset
 | 
240  | 
fun is_dummy thm =  | 
| 
 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 
wenzelm 
parents: 
23599 
diff
changeset
 | 
241  | 
(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
 | 
242  | 
NONE => false  | 
| 
58001
 
934d85f14d1d
more general dummy: may contain "parked arguments", for example;
 
wenzelm 
parents: 
56245 
diff
changeset
 | 
243  | 
| SOME t => Term.is_dummy_pattern (Term.head_of t));  | 
| 
24048
 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 
wenzelm 
parents: 
23599 
diff
changeset
 | 
244  | 
|
| 
22695
 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 
wenzelm 
parents: 
22682 
diff
changeset
 | 
245  | 
fun plain_prop_of raw_thm =  | 
| 
 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 
wenzelm 
parents: 
22682 
diff
changeset
 | 
246  | 
let  | 
| 
 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 
wenzelm 
parents: 
22682 
diff
changeset
 | 
247  | 
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
 | 
248  | 
    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
 | 
249  | 
in  | 
| 61039 | 250  | 
if not (null (Thm.hyps_of thm)) then  | 
| 
22695
 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 
wenzelm 
parents: 
22682 
diff
changeset
 | 
251  | 
err "theorem may not contain hypotheses"  | 
| 
 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 
wenzelm 
parents: 
22682 
diff
changeset
 | 
252  | 
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
 | 
253  | 
err "theorem may not contain sort hypotheses"  | 
| 61039 | 254  | 
else if not (null (Thm.tpairs_of thm)) then  | 
| 
22695
 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 
wenzelm 
parents: 
22682 
diff
changeset
 | 
255  | 
err "theorem may not contain flex-flex pairs"  | 
| 61039 | 256  | 
else Thm.prop_of thm  | 
| 
22695
 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 
wenzelm 
parents: 
22682 
diff
changeset
 | 
257  | 
end;  | 
| 
 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 
wenzelm 
parents: 
22682 
diff
changeset
 | 
258  | 
|
| 
 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 
wenzelm 
parents: 
22682 
diff
changeset
 | 
259  | 
|
| 30564 | 260  | 
(* 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
 | 
261  | 
|
| 
 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 
wenzelm 
parents: 
23599 
diff
changeset
 | 
262  | 
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
 | 
263  | 
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
 | 
264  | 
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
 | 
265  | 
|
| 33453 | 266  | 
val full_rules = Item_Net.init eq_thm_prop (single o Thm.full_prop_of);  | 
| 33373 | 267  | 
val intro_rules = Item_Net.init eq_thm_prop (single o Thm.concl_of);  | 
268  | 
val elim_rules = Item_Net.init eq_thm_prop (single o Thm.major_prem_of);  | 
|
| 30560 | 269  | 
|
270  | 
||
| 22682 | 271  | 
|
| 61508 | 272  | 
(** declared hyps and sort hyps **)  | 
| 
54984
 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 
wenzelm 
parents: 
53206 
diff
changeset
 | 
273  | 
|
| 
 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 
wenzelm 
parents: 
53206 
diff
changeset
 | 
274  | 
structure Hyps = Proof_Data  | 
| 
 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 
wenzelm 
parents: 
53206 
diff
changeset
 | 
275  | 
(  | 
| 61508 | 276  | 
  type T = {checked_hyps: bool, hyps: Termtab.set, shyps: sort Ord_List.T};
 | 
277  | 
  fun init _ : T = {checked_hyps = true, hyps = Termtab.empty, shyps = []};
 | 
|
| 
54984
 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 
wenzelm 
parents: 
53206 
diff
changeset
 | 
278  | 
);  | 
| 
 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 
wenzelm 
parents: 
53206 
diff
changeset
 | 
279  | 
|
| 61508 | 280  | 
fun map_hyps f = Hyps.map (fn {checked_hyps, hyps, shyps} =>
 | 
281  | 
let val (checked_hyps', hyps', shyps') = f (checked_hyps, hyps, shyps)  | 
|
282  | 
  in {checked_hyps = checked_hyps', hyps = hyps', shyps = shyps'} end);
 | 
|
283  | 
||
284  | 
||
285  | 
(* hyps *)  | 
|
286  | 
||
287  | 
fun declare_hyps raw_ct ctxt = ctxt |> map_hyps (fn (checked_hyps, hyps, shyps) =>  | 
|
288  | 
let  | 
|
289  | 
val ct = Thm.transfer_cterm (Proof_Context.theory_of ctxt) raw_ct;  | 
|
290  | 
val hyps' = Termtab.update (Thm.term_of ct, ()) hyps;  | 
|
291  | 
in (checked_hyps, hyps', shyps) end);  | 
|
| 
54984
 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 
wenzelm 
parents: 
53206 
diff
changeset
 | 
292  | 
|
| 
 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 
wenzelm 
parents: 
53206 
diff
changeset
 | 
293  | 
fun assume_hyps ct ctxt = (Thm.assume ct, declare_hyps ct ctxt);  | 
| 
 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 
wenzelm 
parents: 
53206 
diff
changeset
 | 
294  | 
|
| 61508 | 295  | 
val unchecked_hyps = map_hyps (fn (_, hyps, shyps) => (false, hyps, shyps));  | 
296  | 
||
297  | 
fun restore_hyps ctxt =  | 
|
298  | 
map_hyps (fn (_, hyps, shyps) => (#checked_hyps (Hyps.get ctxt), hyps, shyps));  | 
|
| 
54993
 
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
 
wenzelm 
parents: 
54984 
diff
changeset
 | 
299  | 
|
| 55633 | 300  | 
fun undeclared_hyps context th =  | 
301  | 
Thm.hyps_of th  | 
|
302  | 
|> filter_out  | 
|
303  | 
(case context of  | 
|
304  | 
Context.Theory _ => K false  | 
|
305  | 
| Context.Proof ctxt =>  | 
|
306  | 
(case Hyps.get ctxt of  | 
|
| 61508 | 307  | 
          {checked_hyps = false, ...} => K true
 | 
308  | 
        | {hyps, ...} => Termtab.defined hyps));
 | 
|
| 55633 | 309  | 
|
| 
54993
 
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
 
wenzelm 
parents: 
54984 
diff
changeset
 | 
310  | 
fun check_hyps context th =  | 
| 55633 | 311  | 
(case undeclared_hyps context th of  | 
312  | 
[] => th  | 
|
313  | 
| undeclared =>  | 
|
| 61263 | 314  | 
error (Pretty.string_of (Pretty.big_list "Undeclared hyps:"  | 
315  | 
(map (Pretty.item o single o Syntax.pretty_term (Syntax.init_pretty context)) undeclared))));  | 
|
| 
54984
 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 
wenzelm 
parents: 
53206 
diff
changeset
 | 
316  | 
|
| 
 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 
wenzelm 
parents: 
53206 
diff
changeset
 | 
317  | 
|
| 61508 | 318  | 
(* shyps *)  | 
319  | 
||
320  | 
fun declare_term_sorts t =  | 
|
321  | 
map_hyps (fn (checked_hyps, hyps, shyps) =>  | 
|
322  | 
(checked_hyps, hyps, Sorts.insert_term t shyps));  | 
|
323  | 
||
| 
61509
 
358dfae15d83
print thm wrt. local shyps (from full proof context);
 
wenzelm 
parents: 
61508 
diff
changeset
 | 
324  | 
fun extra_shyps' ctxt th =  | 
| 
 
358dfae15d83
print thm wrt. local shyps (from full proof context);
 
wenzelm 
parents: 
61508 
diff
changeset
 | 
325  | 
Sorts.subtract (#shyps (Hyps.get ctxt)) (Thm.extra_shyps th);  | 
| 
 
358dfae15d83
print thm wrt. local shyps (from full proof context);
 
wenzelm 
parents: 
61508 
diff
changeset
 | 
326  | 
|
| 61508 | 327  | 
fun check_shyps ctxt raw_th =  | 
328  | 
let  | 
|
329  | 
val th = Thm.strip_shyps raw_th;  | 
|
| 
61509
 
358dfae15d83
print thm wrt. local shyps (from full proof context);
 
wenzelm 
parents: 
61508 
diff
changeset
 | 
330  | 
val extra_shyps = extra_shyps' ctxt th;  | 
| 61508 | 331  | 
in  | 
| 
61509
 
358dfae15d83
print thm wrt. local shyps (from full proof context);
 
wenzelm 
parents: 
61508 
diff
changeset
 | 
332  | 
if null extra_shyps then th  | 
| 61508 | 333  | 
else error (Pretty.string_of (Pretty.block (Pretty.str "Pending sort hypotheses:" ::  | 
| 
61509
 
358dfae15d83
print thm wrt. local shyps (from full proof context);
 
wenzelm 
parents: 
61508 
diff
changeset
 | 
334  | 
Pretty.brk 1 :: Pretty.commas (map (Syntax.pretty_sort ctxt) extra_shyps))))  | 
| 61508 | 335  | 
end;  | 
336  | 
||
337  | 
val weaken_sorts' = Thm.weaken_sorts o #shyps o Hyps.get;  | 
|
338  | 
||
339  | 
||
| 
54984
 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 
wenzelm 
parents: 
53206 
diff
changeset
 | 
340  | 
|
| 
24980
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
341  | 
(** basic derived rules **)  | 
| 
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
342  | 
|
| 
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
343  | 
(*Elimination of implication  | 
| 
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
344  | 
A A ==> B  | 
| 
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
345  | 
------------  | 
| 
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
346  | 
B  | 
| 
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
347  | 
*)  | 
| 
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
348  | 
fun elim_implies thA thAB = Thm.implies_elim thAB thA;  | 
| 
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
349  | 
|
| 26653 | 350  | 
|
| 61339 | 351  | 
(* forall_intr_name *)  | 
352  | 
||
353  | 
fun forall_intr_name (a, x) th =  | 
|
354  | 
let  | 
|
355  | 
val th' = Thm.forall_intr x th;  | 
|
356  | 
val prop' = (case Thm.prop_of th' of all $ Abs (_, T, b) => all $ Abs (a, T, b));  | 
|
357  | 
in Thm.renamed_prop prop' th' end;  | 
|
358  | 
||
359  | 
||
| 26653 | 360  | 
(* forall_elim_var(s) *)  | 
361  | 
||
362  | 
local  | 
|
363  | 
||
| 
60951
 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60950 
diff
changeset
 | 
364  | 
fun dest_all ct =  | 
| 
 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60950 
diff
changeset
 | 
365  | 
(case Thm.term_of ct of  | 
| 
 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60950 
diff
changeset
 | 
366  | 
    Const ("Pure.all", _) $ Abs (a, _, _) =>
 | 
| 
 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60950 
diff
changeset
 | 
367  | 
let val (x, ct') = Thm.dest_abs NONE (Thm.dest_arg ct)  | 
| 
 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60950 
diff
changeset
 | 
368  | 
in SOME ((a, Thm.ctyp_of_cterm x), ct') end  | 
| 
 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60950 
diff
changeset
 | 
369  | 
| _ => NONE);  | 
| 
 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60950 
diff
changeset
 | 
370  | 
|
| 
 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60950 
diff
changeset
 | 
371  | 
fun dest_all_list ct =  | 
| 
 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60950 
diff
changeset
 | 
372  | 
(case dest_all ct of  | 
| 
 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60950 
diff
changeset
 | 
373  | 
NONE => []  | 
| 
 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60950 
diff
changeset
 | 
374  | 
| SOME (v, ct') => v :: dest_all_list ct');  | 
| 
 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60950 
diff
changeset
 | 
375  | 
|
| 
 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60950 
diff
changeset
 | 
376  | 
fun forall_elim_vars_list vars i th =  | 
| 26653 | 377  | 
let  | 
| 60950 | 378  | 
val used =  | 
379  | 
(Thm.fold_terms o Term.fold_aterms)  | 
|
380  | 
(fn Var ((x, j), _) => if i = j then insert (op =) x else I | _ => I) th [];  | 
|
| 
60951
 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60950 
diff
changeset
 | 
381  | 
val vars' = (Name.variant_list used (map #1 vars), vars)  | 
| 
 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60950 
diff
changeset
 | 
382  | 
|> ListPair.map (fn (x, (_, T)) => Thm.var ((x, i), T));  | 
| 
 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60950 
diff
changeset
 | 
383  | 
in fold Thm.forall_elim vars' th end;  | 
| 26653 | 384  | 
|
385  | 
in  | 
|
386  | 
||
| 60950 | 387  | 
fun forall_elim_vars i th =  | 
| 
60951
 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60950 
diff
changeset
 | 
388  | 
forall_elim_vars_list (dest_all_list (Thm.cprop_of th)) i th;  | 
| 26653 | 389  | 
|
| 33697 | 390  | 
fun forall_elim_var i th =  | 
| 60950 | 391  | 
let  | 
392  | 
val vars =  | 
|
| 
60951
 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60950 
diff
changeset
 | 
393  | 
(case dest_all (Thm.cprop_of th) of  | 
| 
 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60950 
diff
changeset
 | 
394  | 
SOME (v, _) => [v]  | 
| 
 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60950 
diff
changeset
 | 
395  | 
      | NONE => raise THM ("forall_elim_var", i, [th]));
 | 
| 
 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 
wenzelm 
parents: 
60950 
diff
changeset
 | 
396  | 
in forall_elim_vars_list vars i th end;  | 
| 26653 | 397  | 
|
398  | 
end;  | 
|
399  | 
||
400  | 
||
| 60801 | 401  | 
(* instantiate by left-to-right occurrence of variables *)  | 
402  | 
||
403  | 
fun instantiate' cTs cts thm =  | 
|
404  | 
let  | 
|
405  | 
fun err msg =  | 
|
406  | 
      raise TYPE ("instantiate': " ^ msg,
 | 
|
407  | 
map_filter (Option.map Thm.typ_of) cTs,  | 
|
408  | 
map_filter (Option.map Thm.term_of) cts);  | 
|
409  | 
||
410  | 
fun zip_vars xs ys =  | 
|
411  | 
zip_options xs ys handle ListPair.UnequalLengths =>  | 
|
412  | 
err "more instantiations than variables in thm";  | 
|
413  | 
||
414  | 
val thm' =  | 
|
415  | 
Thm.instantiate ((zip_vars (rev (Thm.fold_terms Term.add_tvars thm [])) cTs), []) thm;  | 
|
416  | 
val thm'' =  | 
|
417  | 
Thm.instantiate ([], zip_vars (rev (Thm.fold_terms Term.add_vars thm' [])) cts) thm';  | 
|
418  | 
in thm'' end;  | 
|
419  | 
||
420  | 
||
| 
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
 | 
421  | 
(* 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
 | 
422  | 
|
| 
 
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
 | 
423  | 
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
 | 
424  | 
let  | 
| 61041 | 425  | 
val fixed =  | 
426  | 
fold Term.add_frees (Thm.terms_of_tpairs (Thm.tpairs_of th) @ Thm.hyps_of th) [];  | 
|
| 60821 | 427  | 
val frees =  | 
428  | 
Thm.fold_atomic_cterms (fn a =>  | 
|
429  | 
(case Thm.term_of a of  | 
|
430  | 
Free v => not (member (op =) fixed v) ? insert (op aconvc) a  | 
|
431  | 
| _ => I)) th [];  | 
|
432  | 
in fold Thm.forall_intr frees th end;  | 
|
| 
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
 | 
433  | 
|
| 
 
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
 | 
434  | 
|
| 
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
 | 
435  | 
(* unvarify_global: global schematic variables *)  | 
| 26653 | 436  | 
|
| 60825 | 437  | 
fun unvarify_global thy th =  | 
| 
24980
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
438  | 
let  | 
| 
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
439  | 
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
 | 
440  | 
val _ = map Logic.unvarify_global (prop :: Thm.hyps_of th)  | 
| 
24980
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
441  | 
handle TERM (msg, _) => raise THM (msg, 0, [th]);  | 
| 
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
442  | 
|
| 32279 | 443  | 
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
 | 
444  | 
val inst = rev (Term.add_vars prop []) |> map (fn ((a, i), T) =>  | 
| 32279 | 445  | 
let val T' = Term_Subst.instantiateT instT T  | 
| 60805 | 446  | 
in (((a, i), T'), Thm.global_cterm_of thy (Free ((a, T')))) end);  | 
447  | 
in Thm.instantiate (map (apsnd (Thm.global_ctyp_of thy)) instT, inst) th end;  | 
|
| 
24980
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
448  | 
|
| 60825 | 449  | 
fun unvarify_axiom thy = unvarify_global thy o Thm.axiom thy;  | 
450  | 
||
| 26653 | 451  | 
|
452  | 
(* close_derivation *)  | 
|
453  | 
||
| 
26628
 
63306cb94313
replaced Drule.close_derivation/Goal.close_result by Thm.close_derivation (removed obsolete compression);
 
wenzelm 
parents: 
25518 
diff
changeset
 | 
454  | 
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
 | 
455  | 
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
 | 
456  | 
else thm;  | 
| 
 
63306cb94313
replaced Drule.close_derivation/Goal.close_result by Thm.close_derivation (removed obsolete compression);
 
wenzelm 
parents: 
25518 
diff
changeset
 | 
457  | 
|
| 
24980
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
458  | 
|
| 
59969
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
459  | 
(* user renaming of parameters in a subgoal *)  | 
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
460  | 
|
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
461  | 
(*The names, if distinct, are used for the innermost parameters of subgoal i;  | 
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
462  | 
preceding parameters may be renamed to make all parameters distinct.*)  | 
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
463  | 
fun rename_params_rule (names, i) st =  | 
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
464  | 
let  | 
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
465  | 
val (_, Bs, Bi, C) = Thm.dest_state (st, i);  | 
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
466  | 
val params = map #1 (Logic.strip_params Bi);  | 
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
467  | 
val short = length params - length names;  | 
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
468  | 
val names' =  | 
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
469  | 
if short < 0 then error "More names than parameters in subgoal!"  | 
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
470  | 
else Name.variant_list names (take short params) @ names;  | 
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
471  | 
val free_names = Term.fold_aterms (fn Free (x, _) => insert (op =) x | _ => I) Bi [];  | 
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
472  | 
val Bi' = Logic.list_rename_params names' Bi;  | 
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
473  | 
in  | 
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
474  | 
(case duplicates (op =) names of  | 
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
475  | 
      a :: _ => (warning ("Can't rename.  Bound variables not distinct: " ^ a); st)
 | 
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
476  | 
| [] =>  | 
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
477  | 
(case inter (op =) names free_names of  | 
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
478  | 
        a :: _ => (warning ("Can't rename.  Bound/Free variable clash: " ^ a); st)
 | 
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
479  | 
| [] => Thm.renamed_prop (Logic.list_implies (Bs @ [Bi'], C)) st))  | 
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
480  | 
end;  | 
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
481  | 
|
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
482  | 
|
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
483  | 
(* preservation of bound variable names *)  | 
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
484  | 
|
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
485  | 
fun rename_boundvars pat obj th =  | 
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
486  | 
(case Term.rename_abs pat obj (Thm.prop_of th) of  | 
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
487  | 
NONE => th  | 
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
488  | 
| SOME prop' => Thm.renamed_prop prop' th);  | 
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
489  | 
|
| 
 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 
wenzelm 
parents: 
59623 
diff
changeset
 | 
490  | 
|
| 
24980
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
491  | 
|
| 
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
492  | 
(** specification primitives **)  | 
| 
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
493  | 
|
| 30342 | 494  | 
(* rules *)  | 
495  | 
||
| 
35855
 
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
 
wenzelm 
parents: 
35853 
diff
changeset
 | 
496  | 
fun stripped_sorts thy t =  | 
| 
 
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
 
wenzelm 
parents: 
35853 
diff
changeset
 | 
497  | 
let  | 
| 
60642
 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 
wenzelm 
parents: 
60367 
diff
changeset
 | 
498  | 
val tfrees = rev (Term.add_tfrees t []);  | 
| 
 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 
wenzelm 
parents: 
60367 
diff
changeset
 | 
499  | 
val tfrees' = map (fn a => (a, [])) (Name.invent Name.context Name.aT (length tfrees));  | 
| 
 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 
wenzelm 
parents: 
60367 
diff
changeset
 | 
500  | 
val recover =  | 
| 
 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 
wenzelm 
parents: 
60367 
diff
changeset
 | 
501  | 
map2 (fn (a', S') => fn (a, S) => (((a', 0), S'), Thm.global_ctyp_of thy (TVar ((a, 0), S))))  | 
| 
 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 
wenzelm 
parents: 
60367 
diff
changeset
 | 
502  | 
tfrees' tfrees;  | 
| 
 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 
wenzelm 
parents: 
60367 
diff
changeset
 | 
503  | 
val strip = map (apply2 TFree) (tfrees ~~ tfrees');  | 
| 
35855
 
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
 
wenzelm 
parents: 
35853 
diff
changeset
 | 
504  | 
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
 | 
505  | 
in (strip, recover, t') end;  | 
| 
 
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
 
wenzelm 
parents: 
35853 
diff
changeset
 | 
506  | 
|
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
40238 
diff
changeset
 | 
507  | 
fun add_axiom ctxt (b, prop) thy =  | 
| 
24980
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
508  | 
let  | 
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
40238 
diff
changeset
 | 
509  | 
val _ = Sign.no_vars ctxt prop;  | 
| 
35855
 
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
 
wenzelm 
parents: 
35853 
diff
changeset
 | 
510  | 
val (strip, recover, prop') = stripped_sorts thy prop;  | 
| 
 
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
 
wenzelm 
parents: 
35853 
diff
changeset
 | 
511  | 
val constraints = map (fn (TFree (_, S), T) => (T, S)) strip;  | 
| 60367 | 512  | 
val of_sorts = maps (fn (T as TFree (_, S), _) => of_sort (Thm.ctyp_of ctxt T, S)) strip;  | 
| 
36106
 
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
 
wenzelm 
parents: 
35988 
diff
changeset
 | 
513  | 
|
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
40238 
diff
changeset
 | 
514  | 
val thy' = thy  | 
| 
51316
 
dfe469293eb4
discontinued empty name bindings in 'axiomatization';
 
wenzelm 
parents: 
49062 
diff
changeset
 | 
515  | 
|> Theory.add_axiom ctxt (b, Logic.list_implies (maps Logic.mk_of_sort constraints, prop'));  | 
| 
 
dfe469293eb4
discontinued empty name bindings in 'axiomatization';
 
wenzelm 
parents: 
49062 
diff
changeset
 | 
516  | 
val axm_name = Sign.full_name thy' b;  | 
| 
36106
 
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
 
wenzelm 
parents: 
35988 
diff
changeset
 | 
517  | 
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
 | 
518  | 
val thm =  | 
| 
 
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
 
wenzelm 
parents: 
35985 
diff
changeset
 | 
519  | 
Thm.instantiate (recover, []) axm'  | 
| 60825 | 520  | 
|> unvarify_global thy'  | 
| 
35988
 
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
 
wenzelm 
parents: 
35985 
diff
changeset
 | 
521  | 
|> fold elim_implies of_sorts;  | 
| 
36106
 
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
 
wenzelm 
parents: 
35988 
diff
changeset
 | 
522  | 
in ((axm_name, thm), thy') end;  | 
| 
24980
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
523  | 
|
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
40238 
diff
changeset
 | 
524  | 
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
 | 
525  | 
|
| 
61261
 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 
wenzelm 
parents: 
61059 
diff
changeset
 | 
526  | 
fun add_def (context as (ctxt, _)) unchecked overloaded (b, prop) thy =  | 
| 
24980
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
527  | 
let  | 
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
40238 
diff
changeset
 | 
528  | 
val _ = Sign.no_vars ctxt prop;  | 
| 60367 | 529  | 
val prems = map (Thm.cterm_of ctxt) (Logic.strip_imp_prems prop);  | 
| 
35988
 
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
 
wenzelm 
parents: 
35985 
diff
changeset
 | 
530  | 
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
 | 
531  | 
|
| 
61261
 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 
wenzelm 
parents: 
61059 
diff
changeset
 | 
532  | 
val thy' = Theory.add_def context unchecked overloaded (b, concl') thy;  | 
| 
36106
 
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
 
wenzelm 
parents: 
35988 
diff
changeset
 | 
533  | 
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
 | 
534  | 
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
 | 
535  | 
val thm =  | 
| 
 
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
 
wenzelm 
parents: 
35985 
diff
changeset
 | 
536  | 
Thm.instantiate (recover, []) axm'  | 
| 60825 | 537  | 
|> unvarify_global thy'  | 
| 
35988
 
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
 
wenzelm 
parents: 
35985 
diff
changeset
 | 
538  | 
|> fold_rev Thm.implies_intr prems;  | 
| 
36106
 
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
 
wenzelm 
parents: 
35988 
diff
changeset
 | 
539  | 
in ((axm_name, thm), thy') end;  | 
| 
24980
 
16a74cfca971
added elim_implies (more convenient argument order);
 
wenzelm 
parents: 
24948 
diff
changeset
 | 
540  | 
|
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
40238 
diff
changeset
 | 
541  | 
fun add_def_global unchecked overloaded arg thy =  | 
| 
61262
 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 
wenzelm 
parents: 
61261 
diff
changeset
 | 
542  | 
add_def (Defs.global_context thy) unchecked overloaded arg thy;  | 
| 
42375
 
774df7c59508
report Name_Space.declare/define, relatively to context;
 
wenzelm 
parents: 
40238 
diff
changeset
 | 
543  | 
|
| 
27866
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
544  | 
|
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
545  | 
|
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
546  | 
(*** theorem tags ***)  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
547  | 
|
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
548  | 
(* add / delete tags *)  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
549  | 
|
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
550  | 
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
 | 
551  | 
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
 | 
552  | 
|
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
553  | 
|
| 61852 | 554  | 
(* free dummy thm -- for abstract closure *)  | 
555  | 
||
556  | 
val free_dummyN = "free_dummy";  | 
|
557  | 
fun is_free_dummy thm = Properties.defined (Thm.get_tags thm) free_dummyN;  | 
|
558  | 
val tag_free_dummy = tag_rule (free_dummyN, "");  | 
|
559  | 
||
560  | 
||
| 30342 | 561  | 
(* def_name *)  | 
562  | 
||
563  | 
fun def_name c = c ^ "_def";  | 
|
564  | 
||
565  | 
fun def_name_optional c "" = def_name c  | 
|
566  | 
| def_name_optional _ name = name;  | 
|
567  | 
||
| 
63041
 
cb495c4807b3
clarified def binding position: reset for implicit/derived binding, keep for explicit binding;
 
wenzelm 
parents: 
62093 
diff
changeset
 | 
568  | 
val def_binding = Binding.map_name def_name #> Binding.reset_pos;  | 
| 
 
cb495c4807b3
clarified def binding position: reset for implicit/derived binding, keep for explicit binding;
 
wenzelm 
parents: 
62093 
diff
changeset
 | 
569  | 
fun def_binding_optional b name = if Binding.is_empty name then def_binding b else name;  | 
| 
 
cb495c4807b3
clarified def binding position: reset for implicit/derived binding, keep for explicit binding;
 
wenzelm 
parents: 
62093 
diff
changeset
 | 
570  | 
fun make_def_binding cond b = if cond then def_binding b else Binding.empty;  | 
| 62093 | 571  | 
|
| 30342 | 572  | 
|
| 
27866
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
573  | 
(* unofficial theorem names *)  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
574  | 
|
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
575  | 
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
 | 
576  | 
|
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
577  | 
val has_name_hint = can the_name_hint;  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
578  | 
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
 | 
579  | 
|
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
580  | 
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
 | 
581  | 
|
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
582  | 
|
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
583  | 
(* theorem kinds *)  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
584  | 
|
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
585  | 
val theoremK = "theorem";  | 
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
586  | 
|
| 42473 | 587  | 
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
 | 
588  | 
|
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
589  | 
fun kind_rule k = tag_rule (Markup.kindN, k) o untag_rule Markup.kindN;  | 
| 
61853
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
590  | 
|
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
591  | 
|
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
592  | 
|
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
593  | 
(** attributes **)  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
594  | 
|
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
595  | 
(*attributes subsume any kind of rules or context modifiers*)  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
596  | 
type attribute = Context.generic * thm -> Context.generic option * thm option;  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
597  | 
|
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
598  | 
type binding = binding * attribute list;  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
599  | 
|
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
600  | 
fun rule_attribute ths f (x, th) =  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
601  | 
(NONE,  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
602  | 
(case find_first is_free_dummy (th :: ths) of  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
603  | 
SOME th' => SOME th'  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
604  | 
| NONE => SOME (f x th)));  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
605  | 
|
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
606  | 
fun declaration_attribute f (x, th) =  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
607  | 
(if is_free_dummy th then NONE else SOME (f th x), NONE);  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
608  | 
|
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
609  | 
fun mixed_attribute f (x, th) =  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
610  | 
let val (x', th') = f (x, th) in (SOME x', SOME th') end;  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
611  | 
|
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
612  | 
fun apply_attribute (att: attribute) th x =  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
613  | 
let val (x', th') = att (x, check_hyps x (Thm.transfer (Context.theory_of x) th))  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
614  | 
in (the_default th th', the_default x x') end;  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
615  | 
|
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
616  | 
fun attribute_declaration att th x = #2 (apply_attribute att th x);  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
617  | 
|
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
618  | 
fun apply_attributes mk dest =  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
619  | 
let  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
620  | 
fun app [] th x = (th, x)  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
621  | 
| app (att :: atts) th x = apply_attribute att th (mk x) ||> dest |-> app atts;  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
622  | 
in app end;  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
623  | 
|
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
624  | 
val theory_attributes = apply_attributes Context.Theory Context.the_theory;  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
625  | 
val proof_attributes = apply_attributes Context.Proof Context.the_proof;  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
626  | 
|
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
627  | 
fun no_attributes x = (x, []);  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
628  | 
fun simple_fact x = [(x, [])];  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
629  | 
|
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
630  | 
fun tag tg = rule_attribute [] (K (tag_rule tg));  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
631  | 
fun untag s = rule_attribute [] (K (untag_rule s));  | 
| 
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61852 
diff
changeset
 | 
632  | 
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
 | 
633  | 
|
| 
 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 
wenzelm 
parents: 
27255 
diff
changeset
 | 
634  | 
|
| 
49011
 
9c68e43502ce
some support for registering forked proofs within Proof.state, using its bottom context;
 
wenzelm 
parents: 
49010 
diff
changeset
 | 
635  | 
(* forked proofs *)  | 
| 49010 | 636  | 
|
| 
49062
 
7e31dfd99ce7
discontinued complicated/unreliable notion of recent proofs within context;
 
wenzelm 
parents: 
49058 
diff
changeset
 | 
637  | 
structure Proofs = Theory_Data  | 
| 49010 | 638  | 
(  | 
| 
49062
 
7e31dfd99ce7
discontinued complicated/unreliable notion of recent proofs within context;
 
wenzelm 
parents: 
49058 
diff
changeset
 | 
639  | 
type T = thm list;  | 
| 
 
7e31dfd99ce7
discontinued complicated/unreliable notion of recent proofs within context;
 
wenzelm 
parents: 
49058 
diff
changeset
 | 
640  | 
val empty = [];  | 
| 49010 | 641  | 
fun extend _ = empty;  | 
642  | 
fun merge _ = empty;  | 
|
643  | 
);  | 
|
644  | 
||
| 61059 | 645  | 
fun register_proofs more_thms =  | 
646  | 
Proofs.map (fold (cons o Thm.trim_context) more_thms);  | 
|
647  | 
||
648  | 
fun join_theory_proofs thy =  | 
|
649  | 
Thm.join_proofs (map (Thm.transfer thy) (rev (Proofs.get thy)));  | 
|
| 49010 | 650  | 
|
651  | 
||
| 61268 | 652  | 
|
653  | 
(** print theorems **)  | 
|
654  | 
||
655  | 
(* options *)  | 
|
656  | 
||
657  | 
val show_consts_raw = Config.declare_option ("show_consts", @{here});
 | 
|
658  | 
val show_consts = Config.bool show_consts_raw;  | 
|
659  | 
||
660  | 
val show_hyps_raw = Config.declare ("show_hyps", @{here}) (fn _ => Config.Bool false);
 | 
|
661  | 
val show_hyps = Config.bool show_hyps_raw;  | 
|
662  | 
||
663  | 
val show_tags_raw = Config.declare ("show_tags", @{here}) (fn _ => Config.Bool false);
 | 
|
664  | 
val show_tags = Config.bool show_tags_raw;  | 
|
665  | 
||
666  | 
||
667  | 
(* pretty_thm etc. *)  | 
|
668  | 
||
669  | 
fun pretty_tag (name, arg) = Pretty.strs [name, quote arg];  | 
|
670  | 
val pretty_tags = Pretty.list "[" "]" o map pretty_tag;  | 
|
671  | 
||
672  | 
fun pretty_flexpair ctxt (t, u) = Pretty.block  | 
|
673  | 
[Syntax.pretty_term ctxt t, Pretty.str " =?=", Pretty.brk 1, Syntax.pretty_term ctxt u];  | 
|
674  | 
||
675  | 
fun pretty_thm_raw ctxt {quote, show_hyps = show_hyps'} raw_th =
 | 
|
676  | 
let  | 
|
677  | 
val show_tags = Config.get ctxt show_tags;  | 
|
678  | 
val show_hyps = Config.get ctxt show_hyps;  | 
|
679  | 
||
680  | 
val th = raw_th  | 
|
681  | 
|> perhaps (try (Thm.transfer (Proof_Context.theory_of ctxt)))  | 
|
682  | 
|> perhaps (try Thm.strip_shyps);  | 
|
683  | 
||
684  | 
val hyps = if show_hyps then Thm.hyps_of th else undeclared_hyps (Context.Proof ctxt) th;  | 
|
| 
61509
 
358dfae15d83
print thm wrt. local shyps (from full proof context);
 
wenzelm 
parents: 
61508 
diff
changeset
 | 
685  | 
val extra_shyps = extra_shyps' ctxt th;  | 
| 61268 | 686  | 
val tags = Thm.get_tags th;  | 
687  | 
val tpairs = Thm.tpairs_of th;  | 
|
688  | 
||
689  | 
val q = if quote then Pretty.quote else I;  | 
|
690  | 
val prt_term = q o Syntax.pretty_term ctxt;  | 
|
691  | 
||
692  | 
||
693  | 
val hlen = length extra_shyps + length hyps + length tpairs;  | 
|
694  | 
val hsymbs =  | 
|
695  | 
if hlen = 0 then []  | 
|
696  | 
else if show_hyps orelse show_hyps' then  | 
|
697  | 
[Pretty.brk 2, Pretty.list "[" "]"  | 
|
698  | 
(map (q o pretty_flexpair ctxt) tpairs @ map prt_term hyps @  | 
|
699  | 
map (Syntax.pretty_sort ctxt) extra_shyps)]  | 
|
700  | 
      else [Pretty.brk 2, Pretty.str ("[" ^ replicate_string hlen "." ^ "]")];
 | 
|
701  | 
val tsymbs =  | 
|
702  | 
if null tags orelse not show_tags then []  | 
|
703  | 
else [Pretty.brk 1, pretty_tags tags];  | 
|
704  | 
in Pretty.block (prt_term (Thm.prop_of th) :: (hsymbs @ tsymbs)) end;  | 
|
705  | 
||
706  | 
fun pretty_thm ctxt = pretty_thm_raw ctxt {quote = false, show_hyps = true};
 | 
|
707  | 
fun pretty_thm_item ctxt th = Pretty.item [pretty_thm ctxt th];  | 
|
708  | 
||
709  | 
fun pretty_thm_global thy =  | 
|
710  | 
  pretty_thm_raw (Syntax.init_pretty_global thy) {quote = false, show_hyps = false};
 | 
|
711  | 
||
712  | 
val string_of_thm = Pretty.string_of oo pretty_thm;  | 
|
713  | 
val string_of_thm_global = Pretty.string_of oo pretty_thm_global;  | 
|
714  | 
||
715  | 
||
| 
22362
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
716  | 
open Thm;  | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
717  | 
|
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
718  | 
end;  | 
| 
 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 
wenzelm 
parents:  
diff
changeset
 | 
719  | 
|
| 32842 | 720  | 
structure Basic_Thm: BASIC_THM = Thm;  | 
721  | 
open Basic_Thm;  |