| author | wenzelm | 
| Tue, 22 Sep 2015 18:06:49 +0200 | |
| changeset 61251 | 2da25a27a616 | 
| parent 61059 | 0306e209fa9e | 
| child 61261 | ddb2da7cb2e4 | 
| 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: 
43780diff
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 | |
| 60952 
762cb38a3147
produce certified vars without access to theory_of_thm, and without context;
 wenzelm parents: 
60951diff
changeset | 23 | val eq_ctyp: ctyp * ctyp -> bool | 
| 24948 | 24 | val aconvc: cterm * cterm -> bool | 
| 60952 
762cb38a3147
produce certified vars without access to theory_of_thm, and without context;
 wenzelm parents: 
60951diff
changeset | 25 | val add_tvars: thm -> ctyp list -> ctyp list | 
| 60818 | 26 | val add_frees: thm -> cterm list -> cterm list | 
| 27 | val add_vars: thm -> cterm list -> cterm list | |
| 60938 | 28 | val all_name: Proof.context -> string * cterm -> cterm -> cterm | 
| 29 | val all: Proof.context -> cterm -> cterm -> cterm | |
| 22907 | 30 | val mk_binop: cterm -> cterm -> cterm -> cterm | 
| 31 | val dest_binop: cterm -> cterm * cterm | |
| 32 | val dest_implies: cterm -> cterm * cterm | |
| 33 | val dest_equals: cterm -> cterm * cterm | |
| 34 | val dest_equals_lhs: cterm -> cterm | |
| 35 | val dest_equals_rhs: cterm -> cterm | |
| 36 | val lhs_of: thm -> cterm | |
| 37 | val rhs_of: thm -> cterm | |
| 22362 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 38 | val thm_ord: thm * thm -> order | 
| 32842 | 39 | val cterm_cache: (cterm -> 'a) -> cterm -> 'a | 
| 40 | val thm_cache: (thm -> 'a) -> thm -> 'a | |
| 23599 | 41 | val is_reflexive: thm -> bool | 
| 22362 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 42 | val eq_thm: thm * thm -> bool | 
| 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 43 | val eq_thm_prop: thm * thm -> bool | 
| 52683 
fb028440473e
more official Thm.eq_thm_strict, without demanding ML equality type;
 wenzelm parents: 
51316diff
changeset | 44 | val eq_thm_strict: thm * thm -> bool | 
| 60817 | 45 | val equiv_thm: theory -> thm * thm -> bool | 
| 31944 | 46 | val class_triv: theory -> class -> thm | 
| 47 | val of_sort: ctyp * sort -> thm list | |
| 60819 | 48 | val check_shyps: Proof.context -> sort list -> thm -> thm | 
| 24048 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
23599diff
changeset | 49 | val is_dummy: thm -> bool | 
| 22695 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 wenzelm parents: 
22682diff
changeset | 50 | val plain_prop_of: thm -> term | 
| 24048 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
23599diff
changeset | 51 | val add_thm: thm -> thm list -> thm list | 
| 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
23599diff
changeset | 52 | val del_thm: thm -> thm list -> thm list | 
| 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
23599diff
changeset | 53 | val merge_thms: thm list * thm list -> thm list | 
| 33453 | 54 | val full_rules: thm Item_Net.T | 
| 30560 | 55 | val intro_rules: thm Item_Net.T | 
| 56 | val elim_rules: thm Item_Net.T | |
| 54984 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 wenzelm parents: 
53206diff
changeset | 57 | val declare_hyps: cterm -> Proof.context -> Proof.context | 
| 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 wenzelm parents: 
53206diff
changeset | 58 | 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: 
54984diff
changeset | 59 | val unchecked_hyps: Proof.context -> Proof.context | 
| 
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
 wenzelm parents: 
54984diff
changeset | 60 | val restore_hyps: Proof.context -> Proof.context -> Proof.context | 
| 55633 | 61 | val undeclared_hyps: Context.generic -> thm -> term list | 
| 54993 
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
 wenzelm parents: 
54984diff
changeset | 62 | val check_hyps: Context.generic -> thm -> thm | 
| 27866 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 63 | val elim_implies: thm -> thm -> thm | 
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 64 | val forall_elim_var: int -> thm -> thm | 
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 65 | val forall_elim_vars: int -> thm -> thm | 
| 60801 | 66 | 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: 
35857diff
changeset | 67 | val forall_intr_frees: thm -> thm | 
| 60825 | 68 | val unvarify_global: theory -> thm -> thm | 
| 69 | val unvarify_axiom: theory -> string -> thm | |
| 27866 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 70 | val close_derivation: thm -> thm | 
| 59969 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 71 | val rename_params_rule: string list * int -> thm -> thm | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 72 | val rename_boundvars: term -> term -> thm -> thm | 
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
40238diff
changeset | 73 | val add_axiom: Proof.context -> binding * term -> theory -> (string * thm) * theory | 
| 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
40238diff
changeset | 74 | val add_axiom_global: binding * term -> theory -> (string * thm) * theory | 
| 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
40238diff
changeset | 75 | val add_def: Proof.context -> bool -> bool -> binding * term -> theory -> (string * thm) * theory | 
| 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
40238diff
changeset | 76 | 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: 
43780diff
changeset | 77 | type attribute = Context.generic * thm -> Context.generic option * thm option | 
| 30210 | 78 | type binding = binding * attribute list | 
| 79 | val empty_binding: binding | |
| 27866 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 80 | val rule_attribute: (Context.generic -> thm -> thm) -> attribute | 
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 81 | 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: 
43780diff
changeset | 82 | val mixed_attribute: (Context.generic * thm -> Context.generic * thm) -> attribute | 
| 46775 
6287653e63ec
canonical argument order for attribute application;
 wenzelm parents: 
46497diff
changeset | 83 | 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: 
43780diff
changeset | 84 | val attribute_declaration: attribute -> thm -> Context.generic -> Context.generic | 
| 46775 
6287653e63ec
canonical argument order for attribute application;
 wenzelm parents: 
46497diff
changeset | 85 | val theory_attributes: attribute list -> thm -> theory -> thm * theory | 
| 
6287653e63ec
canonical argument order for attribute application;
 wenzelm parents: 
46497diff
changeset | 86 | val proof_attributes: attribute list -> thm -> Proof.context -> thm * Proof.context | 
| 27866 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 87 | val no_attributes: 'a -> 'a * 'b list | 
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 88 |   val simple_fact: 'a -> ('a * 'b list) list
 | 
| 46830 | 89 | val tag_rule: string * string -> thm -> thm | 
| 27866 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 90 | val untag_rule: string -> thm -> thm | 
| 46830 | 91 | val tag: string * string -> attribute | 
| 27866 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 92 | val untag: string -> attribute | 
| 30342 | 93 | val def_name: string -> string | 
| 94 | val def_name_optional: string -> string -> string | |
| 35238 | 95 | val def_binding: Binding.binding -> Binding.binding | 
| 30433 
ce5138c92ca7
added def_binding_optional -- robust version of def_name_optional for bindings;
 wenzelm parents: 
30342diff
changeset | 96 | val def_binding_optional: Binding.binding -> Binding.binding -> Binding.binding | 
| 27866 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 97 | val has_name_hint: thm -> bool | 
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 98 | val get_name_hint: thm -> string | 
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 99 | val put_name_hint: string -> thm -> thm | 
| 22362 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 100 | val theoremK: string | 
| 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 101 | val lemmaK: string | 
| 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 102 | val corollaryK: string | 
| 42473 | 103 | val legacy_get_kind: thm -> string | 
| 27866 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 104 | val kind_rule: string -> thm -> thm | 
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 105 | val kind: string -> attribute | 
| 49062 
7e31dfd99ce7
discontinued complicated/unreliable notion of recent proofs within context;
 wenzelm parents: 
49058diff
changeset | 106 | val register_proofs: thm list -> theory -> theory | 
| 49011 
9c68e43502ce
some support for registering forked proofs within Proof.state, using its bottom context;
 wenzelm parents: 
49010diff
changeset | 107 | val join_theory_proofs: theory -> unit | 
| 22362 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 108 | end; | 
| 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 109 | |
| 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 110 | structure Thm: THM = | 
| 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 111 | struct | 
| 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 112 | |
| 22695 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 wenzelm parents: 
22682diff
changeset | 113 | (** basic operations **) | 
| 22362 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 114 | |
| 60952 
762cb38a3147
produce certified vars without access to theory_of_thm, and without context;
 wenzelm parents: 
60951diff
changeset | 115 | (* collecting ctyps and cterms *) | 
| 23491 | 116 | |
| 60952 
762cb38a3147
produce certified vars without access to theory_of_thm, and without context;
 wenzelm parents: 
60951diff
changeset | 117 | val eq_ctyp = op = o apply2 Thm.typ_of; | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58001diff
changeset | 118 | val op aconvc = op aconv o apply2 Thm.term_of; | 
| 23491 | 119 | |
| 60952 
762cb38a3147
produce certified vars without access to theory_of_thm, and without context;
 wenzelm parents: 
60951diff
changeset | 120 | val add_tvars = Thm.fold_atomic_ctyps (fn a => is_TVar (Thm.typ_of a) ? insert eq_ctyp a); | 
| 60818 | 121 | val add_frees = Thm.fold_atomic_cterms (fn a => is_Free (Thm.term_of a) ? insert (op aconvc) a); | 
| 122 | val add_vars = Thm.fold_atomic_cterms (fn a => is_Var (Thm.term_of a) ? insert (op aconvc) a); | |
| 23491 | 123 | |
| 124 | ||
| 22907 | 125 | (* cterm constructors and destructors *) | 
| 126 | ||
| 60938 | 127 | fun all_name ctxt (x, t) A = | 
| 32198 | 128 | let | 
| 59586 | 129 | val T = Thm.typ_of_cterm t; | 
| 60938 | 130 |     val all_const = Thm.cterm_of ctxt (Const ("Pure.all", (T --> propT) --> propT));
 | 
| 131 | in Thm.apply all_const (Thm.lambda_name (x, t) A) end; | |
| 32198 | 132 | |
| 60938 | 133 | fun all ctxt t A = all_name ctxt ("", t) A;
 | 
| 32198 | 134 | |
| 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: 
45382diff
changeset | 135 | fun mk_binop c a b = Thm.apply (Thm.apply c a) b; | 
| 22907 | 136 | fun dest_binop ct = (Thm.dest_arg1 ct, Thm.dest_arg ct); | 
| 137 | ||
| 138 | fun dest_implies ct = | |
| 139 | (case Thm.term_of ct of | |
| 56245 | 140 |     Const ("Pure.imp", _) $ _ $ _ => dest_binop ct
 | 
| 22907 | 141 |   | _ => raise TERM ("dest_implies", [Thm.term_of ct]));
 | 
| 142 | ||
| 143 | fun dest_equals ct = | |
| 144 | (case Thm.term_of ct of | |
| 56245 | 145 |     Const ("Pure.eq", _) $ _ $ _ => dest_binop ct
 | 
| 22907 | 146 |   | _ => raise TERM ("dest_equals", [Thm.term_of ct]));
 | 
| 147 | ||
| 148 | fun dest_equals_lhs ct = | |
| 149 | (case Thm.term_of ct of | |
| 56245 | 150 |     Const ("Pure.eq", _) $ _ $ _ => Thm.dest_arg1 ct
 | 
| 22907 | 151 |   | _ => raise TERM ("dest_equals_lhs", [Thm.term_of ct]));
 | 
| 152 | ||
| 153 | fun dest_equals_rhs ct = | |
| 154 | (case Thm.term_of ct of | |
| 56245 | 155 |     Const ("Pure.eq", _) $ _ $ _ => Thm.dest_arg ct
 | 
| 22907 | 156 |   | _ => raise TERM ("dest_equals_rhs", [Thm.term_of ct]));
 | 
| 157 | ||
| 158 | val lhs_of = dest_equals_lhs o Thm.cprop_of; | |
| 159 | val rhs_of = dest_equals_rhs o Thm.cprop_of; | |
| 160 | ||
| 161 | ||
| 162 | (* thm order: ignores theory context! *) | |
| 22682 | 163 | |
| 61039 | 164 | fun thm_ord ths = | 
| 165 | (case Term_Ord.fast_term_ord (apply2 Thm.prop_of ths) of | |
| 166 | EQUAL => | |
| 167 | (case | |
| 168 | list_ord (prod_ord Term_Ord.fast_term_ord Term_Ord.fast_term_ord) | |
| 169 | (apply2 Thm.tpairs_of ths) | |
| 170 | of | |
| 171 | EQUAL => | |
| 172 | (case list_ord Term_Ord.fast_term_ord (apply2 Thm.hyps_of ths) of | |
| 173 | EQUAL => list_ord Term_Ord.sort_ord (apply2 Thm.shyps_of ths) | |
| 174 | | ord => ord) | |
| 175 | | ord => ord) | |
| 176 | | ord => ord); | |
| 22362 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 177 | |
| 22682 | 178 | |
| 32842 | 179 | (* tables and caches *) | 
| 180 | ||
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58001diff
changeset | 181 | structure Ctermtab = Table(type key = cterm val ord = Term_Ord.fast_term_ord o apply2 Thm.term_of); | 
| 32842 | 182 | structure Thmtab = Table(type key = thm val ord = thm_ord); | 
| 183 | ||
| 184 | fun cterm_cache f = Cache.create Ctermtab.empty Ctermtab.lookup Ctermtab.update f; | |
| 185 | fun thm_cache f = Cache.create Thmtab.empty Thmtab.lookup Thmtab.update f; | |
| 186 | ||
| 187 | ||
| 22682 | 188 | (* equality *) | 
| 189 | ||
| 23599 | 190 | fun is_reflexive th = op aconv (Logic.dest_equals (Thm.prop_of th)) | 
| 191 | handle TERM _ => false; | |
| 192 | ||
| 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: 
54996diff
changeset | 193 | val eq_thm = is_equal o thm_ord; | 
| 22362 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 194 | |
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58001diff
changeset | 195 | 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 | 196 | |
| 52683 
fb028440473e
more official Thm.eq_thm_strict, without demanding ML equality type;
 wenzelm parents: 
51316diff
changeset | 197 | 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: 
54996diff
changeset | 198 | eq_thm ths andalso | 
| 61040 | 199 | Context.eq_thy_id (apply2 Thm.theory_id_of_thm ths) andalso | 
| 200 | op = (apply2 Thm.maxidx_of ths) andalso | |
| 201 | op = (apply2 Thm.get_tags ths); | |
| 52683 
fb028440473e
more official Thm.eq_thm_strict, without demanding ML equality type;
 wenzelm parents: 
51316diff
changeset | 202 | |
| 22682 | 203 | |
| 204 | (* pattern equivalence *) | |
| 205 | ||
| 60817 | 206 | fun equiv_thm thy ths = | 
| 207 | 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 | 208 | |
| 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 209 | |
| 31904 
a86896359ca4
renamed Drule.sort_triv to Thm.sort_triv (cf. more_thm.ML);
 wenzelm parents: 
31177diff
changeset | 210 | (* type classes and sorts *) | 
| 
a86896359ca4
renamed Drule.sort_triv to Thm.sort_triv (cf. more_thm.ML);
 wenzelm parents: 
31177diff
changeset | 211 | |
| 31944 | 212 | fun class_triv thy c = | 
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59616diff
changeset | 213 | Thm.of_class (Thm.global_ctyp_of thy (TVar ((Name.aT, 0), [c])), c); | 
| 31944 | 214 | |
| 215 | 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: 
28116diff
changeset | 216 | |
| 60819 | 217 | fun check_shyps ctxt sorts raw_th = | 
| 28621 
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
 wenzelm parents: 
28116diff
changeset | 218 | let | 
| 
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
 wenzelm parents: 
28116diff
changeset | 219 | val th = Thm.strip_shyps raw_th; | 
| 
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
 wenzelm parents: 
28116diff
changeset | 220 | val pending = Sorts.subtract sorts (Thm.extra_shyps th); | 
| 
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
 wenzelm parents: 
28116diff
changeset | 221 | in | 
| 
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
 wenzelm parents: 
28116diff
changeset | 222 | if null pending then th | 
| 
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
 wenzelm parents: 
28116diff
changeset | 223 | else error (Pretty.string_of (Pretty.block (Pretty.str "Pending sort hypotheses:" :: | 
| 60819 | 224 | Pretty.brk 1 :: Pretty.commas (map (Syntax.pretty_sort ctxt) pending)))) | 
| 28621 
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
 wenzelm parents: 
28116diff
changeset | 225 | end; | 
| 
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
 wenzelm parents: 
28116diff
changeset | 226 | |
| 
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
 wenzelm parents: 
28116diff
changeset | 227 | |
| 22695 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 wenzelm parents: 
22682diff
changeset | 228 | (* misc operations *) | 
| 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 wenzelm parents: 
22682diff
changeset | 229 | |
| 24048 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
23599diff
changeset | 230 | fun is_dummy thm = | 
| 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
23599diff
changeset | 231 | (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: 
23599diff
changeset | 232 | NONE => false | 
| 58001 
934d85f14d1d
more general dummy: may contain "parked arguments", for example;
 wenzelm parents: 
56245diff
changeset | 233 | | 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: 
23599diff
changeset | 234 | |
| 22695 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 wenzelm parents: 
22682diff
changeset | 235 | fun plain_prop_of raw_thm = | 
| 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 wenzelm parents: 
22682diff
changeset | 236 | let | 
| 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 wenzelm parents: 
22682diff
changeset | 237 | val thm = Thm.strip_shyps raw_thm; | 
| 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 wenzelm parents: 
22682diff
changeset | 238 |     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: 
22682diff
changeset | 239 | in | 
| 61039 | 240 | if not (null (Thm.hyps_of thm)) then | 
| 22695 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 wenzelm parents: 
22682diff
changeset | 241 | err "theorem may not contain hypotheses" | 
| 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 wenzelm parents: 
22682diff
changeset | 242 | else if not (null (Thm.extra_shyps thm)) then | 
| 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 wenzelm parents: 
22682diff
changeset | 243 | err "theorem may not contain sort hypotheses" | 
| 61039 | 244 | 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: 
22682diff
changeset | 245 | err "theorem may not contain flex-flex pairs" | 
| 61039 | 246 | else Thm.prop_of thm | 
| 22695 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 wenzelm parents: 
22682diff
changeset | 247 | end; | 
| 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 wenzelm parents: 
22682diff
changeset | 248 | |
| 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 wenzelm parents: 
22682diff
changeset | 249 | |
| 30564 | 250 | (* collections of theorems in canonical order *) | 
| 24048 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
23599diff
changeset | 251 | |
| 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
23599diff
changeset | 252 | val add_thm = update eq_thm_prop; | 
| 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
23599diff
changeset | 253 | val del_thm = remove eq_thm_prop; | 
| 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
23599diff
changeset | 254 | val merge_thms = merge eq_thm_prop; | 
| 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
23599diff
changeset | 255 | |
| 33453 | 256 | val full_rules = Item_Net.init eq_thm_prop (single o Thm.full_prop_of); | 
| 33373 | 257 | val intro_rules = Item_Net.init eq_thm_prop (single o Thm.concl_of); | 
| 258 | val elim_rules = Item_Net.init eq_thm_prop (single o Thm.major_prem_of); | |
| 30560 | 259 | |
| 260 | ||
| 22682 | 261 | |
| 54984 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 wenzelm parents: 
53206diff
changeset | 262 | (** declared hyps **) | 
| 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 wenzelm parents: 
53206diff
changeset | 263 | |
| 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 wenzelm parents: 
53206diff
changeset | 264 | structure Hyps = Proof_Data | 
| 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 wenzelm parents: 
53206diff
changeset | 265 | ( | 
| 54993 
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
 wenzelm parents: 
54984diff
changeset | 266 | type T = Termtab.set * bool; | 
| 
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
 wenzelm parents: 
54984diff
changeset | 267 | fun init _ : T = (Termtab.empty, true); | 
| 54984 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 wenzelm parents: 
53206diff
changeset | 268 | ); | 
| 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 wenzelm parents: 
53206diff
changeset | 269 | |
| 60324 | 270 | fun declare_hyps raw_ct ctxt = | 
| 271 | let val ct = Thm.transfer_cterm (Proof_Context.theory_of ctxt) raw_ct | |
| 272 | in (Hyps.map o apfst) (Termtab.update (Thm.term_of ct, ())) ctxt end; | |
| 54984 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 wenzelm parents: 
53206diff
changeset | 273 | |
| 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 wenzelm parents: 
53206diff
changeset | 274 | fun assume_hyps ct ctxt = (Thm.assume ct, declare_hyps ct ctxt); | 
| 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 wenzelm parents: 
53206diff
changeset | 275 | |
| 54993 
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
 wenzelm parents: 
54984diff
changeset | 276 | val unchecked_hyps = (Hyps.map o apsnd) (K false); | 
| 
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
 wenzelm parents: 
54984diff
changeset | 277 | fun restore_hyps ctxt = (Hyps.map o apsnd) (K (#2 (Hyps.get ctxt))); | 
| 
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
 wenzelm parents: 
54984diff
changeset | 278 | |
| 55633 | 279 | fun undeclared_hyps context th = | 
| 280 | Thm.hyps_of th | |
| 281 | |> filter_out | |
| 282 | (case context of | |
| 283 | Context.Theory _ => K false | |
| 284 | | Context.Proof ctxt => | |
| 285 | (case Hyps.get ctxt of | |
| 286 | (_, false) => K true | |
| 287 | | (hyps, _) => Termtab.defined hyps)); | |
| 288 | ||
| 54993 
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
 wenzelm parents: 
54984diff
changeset | 289 | fun check_hyps context th = | 
| 55633 | 290 | (case undeclared_hyps context th of | 
| 291 | [] => th | |
| 292 | | undeclared => | |
| 54993 
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
 wenzelm parents: 
54984diff
changeset | 293 | let | 
| 
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
 wenzelm parents: 
54984diff
changeset | 294 | val ctxt = Context.cases Syntax.init_pretty_global I context; | 
| 
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
 wenzelm parents: 
54984diff
changeset | 295 | in | 
| 
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
 wenzelm parents: 
54984diff
changeset | 296 | error (Pretty.string_of (Pretty.big_list "Undeclared hyps:" | 
| 
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
 wenzelm parents: 
54984diff
changeset | 297 | (map (Pretty.item o single o Syntax.pretty_term ctxt) undeclared))) | 
| 55633 | 298 | end); | 
| 54984 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 wenzelm parents: 
53206diff
changeset | 299 | |
| 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 wenzelm parents: 
53206diff
changeset | 300 | |
| 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 wenzelm parents: 
53206diff
changeset | 301 | |
| 24980 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 302 | (** basic derived rules **) | 
| 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 303 | |
| 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 304 | (*Elimination of implication | 
| 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 305 | A A ==> B | 
| 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 306 | ------------ | 
| 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 307 | B | 
| 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 308 | *) | 
| 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 309 | fun elim_implies thA thAB = Thm.implies_elim thAB thA; | 
| 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 310 | |
| 26653 | 311 | |
| 312 | (* forall_elim_var(s) *) | |
| 313 | ||
| 314 | local | |
| 315 | ||
| 60951 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 wenzelm parents: 
60950diff
changeset | 316 | fun dest_all ct = | 
| 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 wenzelm parents: 
60950diff
changeset | 317 | (case Thm.term_of ct of | 
| 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 wenzelm parents: 
60950diff
changeset | 318 |     Const ("Pure.all", _) $ Abs (a, _, _) =>
 | 
| 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 wenzelm parents: 
60950diff
changeset | 319 | 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: 
60950diff
changeset | 320 | 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: 
60950diff
changeset | 321 | | _ => NONE); | 
| 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 wenzelm parents: 
60950diff
changeset | 322 | |
| 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 wenzelm parents: 
60950diff
changeset | 323 | fun dest_all_list ct = | 
| 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 wenzelm parents: 
60950diff
changeset | 324 | (case dest_all ct of | 
| 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 wenzelm parents: 
60950diff
changeset | 325 | NONE => [] | 
| 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 wenzelm parents: 
60950diff
changeset | 326 | | SOME (v, ct') => v :: dest_all_list ct'); | 
| 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 wenzelm parents: 
60950diff
changeset | 327 | |
| 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 wenzelm parents: 
60950diff
changeset | 328 | fun forall_elim_vars_list vars i th = | 
| 26653 | 329 | let | 
| 60950 | 330 | val used = | 
| 331 | (Thm.fold_terms o Term.fold_aterms) | |
| 332 | (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: 
60950diff
changeset | 333 | 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: 
60950diff
changeset | 334 | |> 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: 
60950diff
changeset | 335 | in fold Thm.forall_elim vars' th end; | 
| 26653 | 336 | |
| 337 | in | |
| 338 | ||
| 60950 | 339 | fun forall_elim_vars i th = | 
| 60951 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 wenzelm parents: 
60950diff
changeset | 340 | forall_elim_vars_list (dest_all_list (Thm.cprop_of th)) i th; | 
| 26653 | 341 | |
| 33697 | 342 | fun forall_elim_var i th = | 
| 60950 | 343 | let | 
| 344 | val vars = | |
| 60951 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 wenzelm parents: 
60950diff
changeset | 345 | (case dest_all (Thm.cprop_of th) of | 
| 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 wenzelm parents: 
60950diff
changeset | 346 | SOME (v, _) => [v] | 
| 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 wenzelm parents: 
60950diff
changeset | 347 |       | NONE => raise THM ("forall_elim_var", i, [th]));
 | 
| 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 wenzelm parents: 
60950diff
changeset | 348 | in forall_elim_vars_list vars i th end; | 
| 26653 | 349 | |
| 350 | end; | |
| 351 | ||
| 352 | ||
| 60801 | 353 | (* instantiate by left-to-right occurrence of variables *) | 
| 354 | ||
| 355 | fun instantiate' cTs cts thm = | |
| 356 | let | |
| 357 | fun err msg = | |
| 358 |       raise TYPE ("instantiate': " ^ msg,
 | |
| 359 | map_filter (Option.map Thm.typ_of) cTs, | |
| 360 | map_filter (Option.map Thm.term_of) cts); | |
| 361 | ||
| 362 | fun zip_vars xs ys = | |
| 363 | zip_options xs ys handle ListPair.UnequalLengths => | |
| 364 | err "more instantiations than variables in thm"; | |
| 365 | ||
| 366 | val thm' = | |
| 367 | Thm.instantiate ((zip_vars (rev (Thm.fold_terms Term.add_tvars thm [])) cTs), []) thm; | |
| 368 | val thm'' = | |
| 369 | Thm.instantiate ([], zip_vars (rev (Thm.fold_terms Term.add_vars thm' [])) cts) thm'; | |
| 370 | in thm'' end; | |
| 371 | ||
| 372 | ||
| 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: 
35857diff
changeset | 373 | (* 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: 
35857diff
changeset | 374 | |
| 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 wenzelm parents: 
35857diff
changeset | 375 | 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: 
35857diff
changeset | 376 | let | 
| 61041 | 377 | val fixed = | 
| 378 | fold Term.add_frees (Thm.terms_of_tpairs (Thm.tpairs_of th) @ Thm.hyps_of th) []; | |
| 60821 | 379 | val frees = | 
| 380 | Thm.fold_atomic_cterms (fn a => | |
| 381 | (case Thm.term_of a of | |
| 382 | Free v => not (member (op =) fixed v) ? insert (op aconvc) a | |
| 383 | | _ => I)) th []; | |
| 384 | 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: 
35857diff
changeset | 385 | |
| 
0bbf0d2348f9
moved Drule.forall_intr_frees to Thm.forall_intr_frees (in more_thm.ML, which is loaded before pure_thy.ML);
 wenzelm parents: 
35857diff
changeset | 386 | |
| 35845 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 wenzelm parents: 
35715diff
changeset | 387 | (* unvarify_global: global schematic variables *) | 
| 26653 | 388 | |
| 60825 | 389 | fun unvarify_global thy th = | 
| 24980 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 390 | let | 
| 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 391 | 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: 
35715diff
changeset | 392 | val _ = map Logic.unvarify_global (prop :: Thm.hyps_of th) | 
| 24980 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 393 | handle TERM (msg, _) => raise THM (msg, 0, [th]); | 
| 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 394 | |
| 32279 | 395 | 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: 
24948diff
changeset | 396 | val inst = rev (Term.add_vars prop []) |> map (fn ((a, i), T) => | 
| 32279 | 397 | let val T' = Term_Subst.instantiateT instT T | 
| 60805 | 398 | in (((a, i), T'), Thm.global_cterm_of thy (Free ((a, T')))) end); | 
| 399 | 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: 
24948diff
changeset | 400 | |
| 60825 | 401 | fun unvarify_axiom thy = unvarify_global thy o Thm.axiom thy; | 
| 402 | ||
| 26653 | 403 | |
| 404 | (* close_derivation *) | |
| 405 | ||
| 26628 
63306cb94313
replaced Drule.close_derivation/Goal.close_result by Thm.close_derivation (removed obsolete compression);
 wenzelm parents: 
25518diff
changeset | 406 | 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: 
36106diff
changeset | 407 | 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: 
25518diff
changeset | 408 | else thm; | 
| 
63306cb94313
replaced Drule.close_derivation/Goal.close_result by Thm.close_derivation (removed obsolete compression);
 wenzelm parents: 
25518diff
changeset | 409 | |
| 24980 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 410 | |
| 59969 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 411 | (* user renaming of parameters in a subgoal *) | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 412 | |
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 413 | (*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: 
59623diff
changeset | 414 | preceding parameters may be renamed to make all parameters distinct.*) | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 415 | fun rename_params_rule (names, i) st = | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 416 | let | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 417 | val (_, Bs, Bi, C) = Thm.dest_state (st, i); | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 418 | val params = map #1 (Logic.strip_params Bi); | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 419 | val short = length params - length names; | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 420 | val names' = | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 421 | if short < 0 then error "More names than parameters in subgoal!" | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 422 | else Name.variant_list names (take short params) @ names; | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 423 | 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: 
59623diff
changeset | 424 | val Bi' = Logic.list_rename_params names' Bi; | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 425 | in | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 426 | (case duplicates (op =) names of | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 427 |       a :: _ => (warning ("Can't rename.  Bound variables not distinct: " ^ a); st)
 | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 428 | | [] => | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 429 | (case inter (op =) names free_names of | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 430 |         a :: _ => (warning ("Can't rename.  Bound/Free variable clash: " ^ a); st)
 | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 431 | | [] => Thm.renamed_prop (Logic.list_implies (Bs @ [Bi'], C)) st)) | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 432 | end; | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 433 | |
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 434 | |
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 435 | (* preservation of bound variable names *) | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 436 | |
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 437 | fun rename_boundvars pat obj th = | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 438 | (case Term.rename_abs pat obj (Thm.prop_of th) of | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 439 | NONE => th | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 440 | | SOME prop' => Thm.renamed_prop prop' th); | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 441 | |
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 442 | |
| 24980 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 443 | |
| 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 444 | (** specification primitives **) | 
| 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 445 | |
| 30342 | 446 | (* rules *) | 
| 447 | ||
| 35855 
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
 wenzelm parents: 
35853diff
changeset | 448 | fun stripped_sorts thy t = | 
| 
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
 wenzelm parents: 
35853diff
changeset | 449 | 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: 
60367diff
changeset | 450 | 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: 
60367diff
changeset | 451 | 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: 
60367diff
changeset | 452 | 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: 
60367diff
changeset | 453 | 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: 
60367diff
changeset | 454 | 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: 
60367diff
changeset | 455 | val strip = map (apply2 TFree) (tfrees ~~ tfrees'); | 
| 35855 
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
 wenzelm parents: 
35853diff
changeset | 456 | 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: 
35853diff
changeset | 457 | in (strip, recover, t') end; | 
| 
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
 wenzelm parents: 
35853diff
changeset | 458 | |
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
40238diff
changeset | 459 | fun add_axiom ctxt (b, prop) thy = | 
| 24980 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 460 | let | 
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
40238diff
changeset | 461 | val _ = Sign.no_vars ctxt prop; | 
| 35855 
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
 wenzelm parents: 
35853diff
changeset | 462 | val (strip, recover, prop') = stripped_sorts thy prop; | 
| 
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
 wenzelm parents: 
35853diff
changeset | 463 | val constraints = map (fn (TFree (_, S), T) => (T, S)) strip; | 
| 60367 | 464 | 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: 
35988diff
changeset | 465 | |
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
40238diff
changeset | 466 | val thy' = thy | 
| 51316 
dfe469293eb4
discontinued empty name bindings in 'axiomatization';
 wenzelm parents: 
49062diff
changeset | 467 | |> Theory.add_axiom ctxt (b, Logic.list_implies (maps Logic.mk_of_sort constraints, prop')); | 
| 
dfe469293eb4
discontinued empty name bindings in 'axiomatization';
 wenzelm parents: 
49062diff
changeset | 468 | val axm_name = Sign.full_name thy' b; | 
| 36106 
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
 wenzelm parents: 
35988diff
changeset | 469 | val axm' = Thm.axiom thy' axm_name; | 
| 35988 
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
 wenzelm parents: 
35985diff
changeset | 470 | val thm = | 
| 
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
 wenzelm parents: 
35985diff
changeset | 471 | Thm.instantiate (recover, []) axm' | 
| 60825 | 472 | |> unvarify_global thy' | 
| 35988 
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
 wenzelm parents: 
35985diff
changeset | 473 | |> fold elim_implies of_sorts; | 
| 36106 
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
 wenzelm parents: 
35988diff
changeset | 474 | in ((axm_name, thm), thy') end; | 
| 24980 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 475 | |
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
40238diff
changeset | 476 | 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: 
40238diff
changeset | 477 | |
| 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
40238diff
changeset | 478 | fun add_def ctxt unchecked overloaded (b, prop) thy = | 
| 24980 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 479 | let | 
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
40238diff
changeset | 480 | val _ = Sign.no_vars ctxt prop; | 
| 60367 | 481 | 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: 
35985diff
changeset | 482 | 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: 
35988diff
changeset | 483 | |
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
40238diff
changeset | 484 | 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: 
35988diff
changeset | 485 | val axm_name = Sign.full_name thy' b; | 
| 
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
 wenzelm parents: 
35988diff
changeset | 486 | val axm' = Thm.axiom thy' axm_name; | 
| 35988 
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
 wenzelm parents: 
35985diff
changeset | 487 | val thm = | 
| 
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
 wenzelm parents: 
35985diff
changeset | 488 | Thm.instantiate (recover, []) axm' | 
| 60825 | 489 | |> unvarify_global thy' | 
| 35988 
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
 wenzelm parents: 
35985diff
changeset | 490 | |> fold_rev Thm.implies_intr prems; | 
| 36106 
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
 wenzelm parents: 
35988diff
changeset | 491 | in ((axm_name, thm), thy') end; | 
| 24980 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 492 | |
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
40238diff
changeset | 493 | fun add_def_global unchecked overloaded arg thy = | 
| 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
40238diff
changeset | 494 | add_def (Syntax.init_pretty_global thy) unchecked overloaded arg thy; | 
| 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
40238diff
changeset | 495 | |
| 27866 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 496 | |
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 497 | |
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 498 | (** attributes **) | 
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 499 | |
| 40238 
edcdecd55655
type attribute is derived concept outside the kernel;
 wenzelm parents: 
39133diff
changeset | 500 | (*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: 
43780diff
changeset | 501 | type attribute = Context.generic * thm -> Context.generic option * thm option; | 
| 40238 
edcdecd55655
type attribute is derived concept outside the kernel;
 wenzelm parents: 
39133diff
changeset | 502 | |
| 30210 | 503 | type binding = binding * attribute list; | 
| 504 | val empty_binding: binding = (Binding.empty, []); | |
| 505 | ||
| 45375 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 wenzelm parents: 
43780diff
changeset | 506 | 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: 
43780diff
changeset | 507 | 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: 
43780diff
changeset | 508 | 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: 
43780diff
changeset | 509 | |
| 46775 
6287653e63ec
canonical argument order for attribute application;
 wenzelm parents: 
46497diff
changeset | 510 | fun apply_attribute (att: attribute) th x = | 
| 54996 | 511 | let val (x', th') = att (x, check_hyps x (Thm.transfer (Context.theory_of x) th)) | 
| 46775 
6287653e63ec
canonical argument order for attribute application;
 wenzelm parents: 
46497diff
changeset | 512 | 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: 
43780diff
changeset | 513 | |
| 46775 
6287653e63ec
canonical argument order for attribute application;
 wenzelm parents: 
46497diff
changeset | 514 | 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: 
27255diff
changeset | 515 | |
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 516 | fun apply_attributes mk dest = | 
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 517 | let | 
| 46775 
6287653e63ec
canonical argument order for attribute application;
 wenzelm parents: 
46497diff
changeset | 518 | fun app [] th x = (th, x) | 
| 
6287653e63ec
canonical argument order for attribute application;
 wenzelm parents: 
46497diff
changeset | 519 | | 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: 
27255diff
changeset | 520 | in app end; | 
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 521 | |
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 522 | val theory_attributes = apply_attributes Context.Theory Context.the_theory; | 
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 523 | val proof_attributes = apply_attributes Context.Proof Context.the_proof; | 
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 524 | |
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 525 | fun no_attributes x = (x, []); | 
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 526 | fun simple_fact x = [(x, [])]; | 
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 527 | |
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 528 | |
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 529 | |
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 530 | (*** theorem tags ***) | 
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 531 | |
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 532 | (* add / delete tags *) | 
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 533 | |
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 534 | fun tag_rule tg = Thm.map_tags (insert (op =) tg); | 
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 535 | 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: 
27255diff
changeset | 536 | |
| 45375 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 wenzelm parents: 
43780diff
changeset | 537 | fun tag tg = rule_attribute (K (tag_rule tg)); | 
| 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 wenzelm parents: 
43780diff
changeset | 538 | fun untag s = rule_attribute (K (untag_rule s)); | 
| 27866 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 539 | |
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 540 | |
| 30342 | 541 | (* def_name *) | 
| 542 | ||
| 543 | fun def_name c = c ^ "_def"; | |
| 544 | ||
| 545 | fun def_name_optional c "" = def_name c | |
| 546 | | def_name_optional _ name = name; | |
| 547 | ||
| 35238 | 548 | val def_binding = Binding.map_name def_name; | 
| 549 | ||
| 30433 
ce5138c92ca7
added def_binding_optional -- robust version of def_name_optional for bindings;
 wenzelm parents: 
30342diff
changeset | 550 | fun def_binding_optional b name = | 
| 35238 | 551 | 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: 
30342diff
changeset | 552 | |
| 30342 | 553 | |
| 27866 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 554 | (* unofficial theorem names *) | 
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 555 | |
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 556 | 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: 
27255diff
changeset | 557 | |
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 558 | val has_name_hint = can the_name_hint; | 
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 559 | val get_name_hint = the_default "??.unknown" o try the_name_hint; | 
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 560 | |
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 561 | 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: 
27255diff
changeset | 562 | |
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 563 | |
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 564 | (* theorem kinds *) | 
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 565 | |
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 566 | val theoremK = "theorem"; | 
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 567 | val lemmaK = "lemma"; | 
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 568 | val corollaryK = "corollary"; | 
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 569 | |
| 42473 | 570 | 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: 
27255diff
changeset | 571 | |
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 572 | 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: 
43780diff
changeset | 573 | fun kind k = rule_attribute (K (k <> "" ? kind_rule k)); | 
| 27866 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 574 | |
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 575 | |
| 49011 
9c68e43502ce
some support for registering forked proofs within Proof.state, using its bottom context;
 wenzelm parents: 
49010diff
changeset | 576 | (* forked proofs *) | 
| 49010 | 577 | |
| 49062 
7e31dfd99ce7
discontinued complicated/unreliable notion of recent proofs within context;
 wenzelm parents: 
49058diff
changeset | 578 | structure Proofs = Theory_Data | 
| 49010 | 579 | ( | 
| 49062 
7e31dfd99ce7
discontinued complicated/unreliable notion of recent proofs within context;
 wenzelm parents: 
49058diff
changeset | 580 | type T = thm list; | 
| 
7e31dfd99ce7
discontinued complicated/unreliable notion of recent proofs within context;
 wenzelm parents: 
49058diff
changeset | 581 | val empty = []; | 
| 49010 | 582 | fun extend _ = empty; | 
| 583 | fun merge _ = empty; | |
| 584 | ); | |
| 585 | ||
| 61059 | 586 | fun register_proofs more_thms = | 
| 587 | Proofs.map (fold (cons o Thm.trim_context) more_thms); | |
| 588 | ||
| 589 | fun join_theory_proofs thy = | |
| 590 | Thm.join_proofs (map (Thm.transfer thy) (rev (Proofs.get thy))); | |
| 49010 | 591 | |
| 592 | ||
| 22362 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 593 | open Thm; | 
| 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 594 | |
| 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 595 | end; | 
| 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 596 | |
| 32842 | 597 | structure Basic_Thm: BASIC_THM = Thm; | 
| 598 | open Basic_Thm; |