| author | paulson <lp15@cam.ac.uk> | 
| Sun, 14 Apr 2024 18:39:43 +0100 | |
| changeset 80101 | 2ff4cc7fa70a | 
| parent 79436 | 253d86888e84 | 
| child 80306 | c2537860ccf8 | 
| 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 | val aconvc: cterm * cterm -> bool | 
| 45375 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 wenzelm parents: 
43780diff
changeset | 16 | type attribute = Context.generic * thm -> Context.generic option * thm option | 
| 32842 | 17 | end; | 
| 18 | ||
| 22362 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 19 | signature THM = | 
| 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 20 | sig | 
| 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 21 | include THM | 
| 60952 
762cb38a3147
produce certified vars without access to theory_of_thm, and without context;
 wenzelm parents: 
60951diff
changeset | 22 | val eq_ctyp: ctyp * ctyp -> bool | 
| 24948 | 23 | val aconvc: cterm * cterm -> bool | 
| 74266 | 24 | val add_tvars: thm -> ctyp TVars.table -> ctyp TVars.table | 
| 25 | val add_vars: thm -> cterm Vars.table -> cterm Vars.table | |
| 70159 | 26 | val dest_funT: ctyp -> ctyp * ctyp | 
| 70152 | 27 | val strip_type: ctyp -> ctyp list * ctyp | 
| 74566 
8e0f0317e266
ML antiquotations to instantiate types/terms/props;
 wenzelm parents: 
74561diff
changeset | 28 | val instantiate_ctyp: ctyp TVars.table -> ctyp -> ctyp | 
| 60938 | 29 | val all_name: Proof.context -> string * cterm -> cterm -> cterm | 
| 30 | val all: Proof.context -> cterm -> cterm -> cterm | |
| 22907 | 31 | val mk_binop: cterm -> cterm -> cterm -> cterm | 
| 32 | val dest_binop: cterm -> cterm * cterm | |
| 33 | val dest_implies: cterm -> cterm * cterm | |
| 34 | val dest_equals: cterm -> cterm * cterm | |
| 35 | val dest_equals_lhs: cterm -> cterm | |
| 36 | val dest_equals_rhs: cterm -> cterm | |
| 37 | val lhs_of: thm -> cterm | |
| 38 | val rhs_of: thm -> cterm | |
| 23599 | 39 | val is_reflexive: thm -> bool | 
| 22362 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 40 | val eq_thm: thm * thm -> bool | 
| 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 41 | val eq_thm_prop: thm * thm -> bool | 
| 52683 
fb028440473e
more official Thm.eq_thm_strict, without demanding ML equality type;
 wenzelm parents: 
51316diff
changeset | 42 | val eq_thm_strict: thm * thm -> bool | 
| 60817 | 43 | val equiv_thm: theory -> thm * thm -> bool | 
| 31944 | 44 | val class_triv: theory -> class -> thm | 
| 45 | val of_sort: ctyp * sort -> thm list | |
| 24048 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
23599diff
changeset | 46 | val is_dummy: thm -> bool | 
| 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
23599diff
changeset | 47 | 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 | 48 | 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 | 49 | val merge_thms: thm list * thm list -> thm list | 
| 74152 | 50 | val item_net: thm Item_Net.T | 
| 51 | val item_net_intro: thm Item_Net.T | |
| 52 | val item_net_elim: thm Item_Net.T | |
| 54984 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 wenzelm parents: 
53206diff
changeset | 53 | val declare_hyps: cterm -> Proof.context -> Proof.context | 
| 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 wenzelm parents: 
53206diff
changeset | 54 | 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 | 55 | val unchecked_hyps: Proof.context -> Proof.context | 
| 
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
 wenzelm parents: 
54984diff
changeset | 56 | val restore_hyps: Proof.context -> Proof.context -> Proof.context | 
| 55633 | 57 | 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 | 58 | val check_hyps: Context.generic -> thm -> thm | 
| 61508 | 59 | val declare_term_sorts: term -> Proof.context -> Proof.context | 
| 77869 | 60 | val extra_shyps': Proof.context -> thm -> sort list | 
| 61508 | 61 | val check_shyps: Proof.context -> thm -> thm | 
| 62 | val weaken_sorts': Proof.context -> cterm -> cterm | |
| 27866 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 63 | val elim_implies: thm -> thm -> thm | 
| 61339 | 64 | val forall_intr_name: string * cterm -> thm -> thm | 
| 27866 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 65 | val forall_elim_var: int -> thm -> thm | 
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 66 | val forall_elim_vars: int -> thm -> thm | 
| 74282 | 67 | val instantiate_frees: ctyp TFrees.table * cterm Frees.table -> thm -> thm | 
| 60801 | 68 | 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 | 69 | val forall_intr_frees: thm -> thm | 
| 74245 | 70 | val forall_intr_vars: thm -> thm | 
| 60825 | 71 | val unvarify_global: theory -> thm -> thm | 
| 72 | val unvarify_axiom: theory -> string -> thm | |
| 59969 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 73 | val rename_params_rule: string list * int -> thm -> thm | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 74 | val rename_boundvars: term -> term -> thm -> thm | 
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
40238diff
changeset | 75 | val add_axiom: Proof.context -> binding * term -> theory -> (string * thm) * theory | 
| 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
40238diff
changeset | 76 | 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: 
61059diff
changeset | 77 | 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: 
40238diff
changeset | 78 | 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 | 79 | type attribute = Context.generic * thm -> Context.generic option * thm option | 
| 30210 | 80 | type binding = binding * attribute list | 
| 46830 | 81 | val tag_rule: string * string -> thm -> thm | 
| 27866 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 82 | val untag_rule: string -> thm -> thm | 
| 61852 | 83 | val is_free_dummy: thm -> bool | 
| 84 | val tag_free_dummy: thm -> thm | |
| 30342 | 85 | val def_name: string -> string | 
| 86 | val def_name_optional: string -> string -> string | |
| 35238 | 87 | val def_binding: Binding.binding -> Binding.binding | 
| 30433 
ce5138c92ca7
added def_binding_optional -- robust version of def_name_optional for bindings;
 wenzelm parents: 
30342diff
changeset | 88 | val def_binding_optional: Binding.binding -> Binding.binding -> Binding.binding | 
| 62093 | 89 | val make_def_binding: bool -> Binding.binding -> Binding.binding | 
| 27866 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 90 | val has_name_hint: thm -> bool | 
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 91 | val get_name_hint: thm -> string | 
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 92 | val put_name_hint: string -> thm -> thm | 
| 22362 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 93 | val theoremK: string | 
| 42473 | 94 | val legacy_get_kind: thm -> string | 
| 27866 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 95 | 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: 
61852diff
changeset | 96 | 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: 
61852diff
changeset | 97 | 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: 
61852diff
changeset | 98 | 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: 
61852diff
changeset | 99 | 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: 
61852diff
changeset | 100 | 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: 
61852diff
changeset | 101 | 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: 
61852diff
changeset | 102 | 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: 
61852diff
changeset | 103 | 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: 
61852diff
changeset | 104 |   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: 
61852diff
changeset | 105 | 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: 
61852diff
changeset | 106 | val untag: string -> attribute | 
| 27866 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 107 | val kind: string -> attribute | 
| 67671 
857da80611ab
support for lazy notes in global/local context and Element.Lazy_Notes: name binding and fact without attributes;
 wenzelm parents: 
67649diff
changeset | 108 | val register_proofs: thm list lazy -> theory -> theory | 
| 64574 
1134e4d5e5b7
consolidate nested thms with persistent result, for improved performance;
 wenzelm parents: 
64568diff
changeset | 109 | val consolidate_theory: theory -> unit | 
| 71023 
35a8e15b7e03
more robust Thm.expose_theory -- ensure that PIDE export happens in the proper theory context;
 wenzelm parents: 
71018diff
changeset | 110 | val expose_theory: theory -> unit | 
| 61268 | 111 | val show_consts: bool Config.T | 
| 112 | val show_hyps: bool Config.T | |
| 113 | val show_tags: bool Config.T | |
| 114 |   val pretty_thm_raw: Proof.context -> {quote: bool, show_hyps: bool} -> thm -> Pretty.T
 | |
| 115 | val pretty_thm: Proof.context -> thm -> Pretty.T | |
| 116 | val pretty_thm_item: Proof.context -> thm -> Pretty.T | |
| 117 | val pretty_thm_global: theory -> thm -> Pretty.T | |
| 118 | val string_of_thm: Proof.context -> thm -> string | |
| 119 | val string_of_thm_global: theory -> thm -> string | |
| 22362 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 120 | end; | 
| 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 121 | |
| 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 122 | structure Thm: THM = | 
| 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 123 | struct | 
| 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 124 | |
| 22695 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 wenzelm parents: 
22682diff
changeset | 125 | (** basic operations **) | 
| 22362 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 126 | |
| 60952 
762cb38a3147
produce certified vars without access to theory_of_thm, and without context;
 wenzelm parents: 
60951diff
changeset | 127 | (* collecting ctyps and cterms *) | 
| 23491 | 128 | |
| 60952 
762cb38a3147
produce certified vars without access to theory_of_thm, and without context;
 wenzelm parents: 
60951diff
changeset | 129 | val eq_ctyp = op = o apply2 Thm.typ_of; | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58001diff
changeset | 130 | val op aconvc = op aconv o apply2 Thm.term_of; | 
| 23491 | 131 | |
| 74241 
eb265f54e3ce
more efficient operations: traverse hyps only when required;
 wenzelm parents: 
74239diff
changeset | 132 | val add_tvars = | 
| 74244 | 133 |   Thm.fold_atomic_ctyps {hyps = false} Term.is_TVar (fn cT => fn tab =>
 | 
| 134 | let val v = Term.dest_TVar (Thm.typ_of cT) | |
| 74266 | 135 | in tab |> not (TVars.defined tab v) ? TVars.add (v, cT) end); | 
| 74243 | 136 | |
| 74241 
eb265f54e3ce
more efficient operations: traverse hyps only when required;
 wenzelm parents: 
74239diff
changeset | 137 | val add_vars = | 
| 74244 | 138 |   Thm.fold_atomic_cterms {hyps = false} Term.is_Var (fn ct => fn tab =>
 | 
| 139 | let val v = Term.dest_Var (Thm.term_of ct) | |
| 74266 | 140 | in tab |> not (Vars.defined tab v) ? Vars.add (v, ct) end); | 
| 23491 | 141 | |
| 142 | ||
| 70155 | 143 | (* ctyp operations *) | 
| 70152 | 144 | |
| 70159 | 145 | fun dest_funT cT = | 
| 70155 | 146 | (case Thm.typ_of cT of | 
| 147 |     Type ("fun", _) => let val [A, B] = Thm.dest_ctyp cT in (A, B) end
 | |
| 70159 | 148 |   | T => raise TYPE ("dest_funT", [T], []));
 | 
| 70152 | 149 | |
| 150 | (* ctyp version of strip_type: maps [T1,...,Tn]--->T to ([T1,T2,...,Tn], T) *) | |
| 151 | fun strip_type cT = | |
| 152 | (case Thm.typ_of cT of | |
| 153 |     Type ("fun", _) =>
 | |
| 154 | let | |
| 70159 | 155 | val (cT1, cT2) = dest_funT cT; | 
| 70152 | 156 | val (cTs, cT') = strip_type cT2 | 
| 157 | in (cT1 :: cTs, cT') end | |
| 158 | | _ => ([], cT)); | |
| 159 | ||
| 74566 
8e0f0317e266
ML antiquotations to instantiate types/terms/props;
 wenzelm parents: 
74561diff
changeset | 160 | fun instantiate_ctyp instT cT = | 
| 
8e0f0317e266
ML antiquotations to instantiate types/terms/props;
 wenzelm parents: 
74561diff
changeset | 161 |   Thm.instantiate_cterm (instT, Vars.empty) (Thm.var (("x", 0), cT))
 | 
| 
8e0f0317e266
ML antiquotations to instantiate types/terms/props;
 wenzelm parents: 
74561diff
changeset | 162 | |> Thm.ctyp_of_cterm; | 
| 
8e0f0317e266
ML antiquotations to instantiate types/terms/props;
 wenzelm parents: 
74561diff
changeset | 163 | |
| 70152 | 164 | |
| 70155 | 165 | (* cterm operations *) | 
| 22907 | 166 | |
| 60938 | 167 | fun all_name ctxt (x, t) A = | 
| 32198 | 168 | let | 
| 59586 | 169 | val T = Thm.typ_of_cterm t; | 
| 60938 | 170 |     val all_const = Thm.cterm_of ctxt (Const ("Pure.all", (T --> propT) --> propT));
 | 
| 171 | in Thm.apply all_const (Thm.lambda_name (x, t) A) end; | |
| 32198 | 172 | |
| 60938 | 173 | fun all ctxt t A = all_name ctxt ("", t) A;
 | 
| 32198 | 174 | |
| 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 | 175 | fun mk_binop c a b = Thm.apply (Thm.apply c a) b; | 
| 22907 | 176 | fun dest_binop ct = (Thm.dest_arg1 ct, Thm.dest_arg ct); | 
| 177 | ||
| 178 | fun dest_implies ct = | |
| 179 | (case Thm.term_of ct of | |
| 56245 | 180 |     Const ("Pure.imp", _) $ _ $ _ => dest_binop ct
 | 
| 22907 | 181 |   | _ => raise TERM ("dest_implies", [Thm.term_of ct]));
 | 
| 182 | ||
| 183 | fun dest_equals ct = | |
| 184 | (case Thm.term_of ct of | |
| 56245 | 185 |     Const ("Pure.eq", _) $ _ $ _ => dest_binop ct
 | 
| 22907 | 186 |   | _ => raise TERM ("dest_equals", [Thm.term_of ct]));
 | 
| 187 | ||
| 188 | fun dest_equals_lhs ct = | |
| 189 | (case Thm.term_of ct of | |
| 56245 | 190 |     Const ("Pure.eq", _) $ _ $ _ => Thm.dest_arg1 ct
 | 
| 22907 | 191 |   | _ => raise TERM ("dest_equals_lhs", [Thm.term_of ct]));
 | 
| 192 | ||
| 193 | fun dest_equals_rhs ct = | |
| 194 | (case Thm.term_of ct of | |
| 56245 | 195 |     Const ("Pure.eq", _) $ _ $ _ => Thm.dest_arg ct
 | 
| 22907 | 196 |   | _ => raise TERM ("dest_equals_rhs", [Thm.term_of ct]));
 | 
| 197 | ||
| 198 | val lhs_of = dest_equals_lhs o Thm.cprop_of; | |
| 199 | val rhs_of = dest_equals_rhs o Thm.cprop_of; | |
| 200 | ||
| 201 | ||
| 22682 | 202 | (* equality *) | 
| 203 | ||
| 23599 | 204 | fun is_reflexive th = op aconv (Logic.dest_equals (Thm.prop_of th)) | 
| 205 | handle TERM _ => false; | |
| 206 | ||
| 74270 | 207 | val eq_thm = is_equal o Thm.thm_ord; | 
| 22362 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 208 | |
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58001diff
changeset | 209 | 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 | 210 | |
| 52683 
fb028440473e
more official Thm.eq_thm_strict, without demanding ML equality type;
 wenzelm parents: 
51316diff
changeset | 211 | 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 | 212 | eq_thm ths andalso | 
| 65458 | 213 | Context.eq_thy_id (apply2 Thm.theory_id ths) andalso | 
| 61040 | 214 | op = (apply2 Thm.maxidx_of ths) andalso | 
| 215 | op = (apply2 Thm.get_tags ths); | |
| 52683 
fb028440473e
more official Thm.eq_thm_strict, without demanding ML equality type;
 wenzelm parents: 
51316diff
changeset | 216 | |
| 22682 | 217 | |
| 218 | (* pattern equivalence *) | |
| 219 | ||
| 60817 | 220 | fun equiv_thm thy ths = | 
| 221 | 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 | 222 | |
| 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 223 | |
| 31904 
a86896359ca4
renamed Drule.sort_triv to Thm.sort_triv (cf. more_thm.ML);
 wenzelm parents: 
31177diff
changeset | 224 | (* type classes and sorts *) | 
| 
a86896359ca4
renamed Drule.sort_triv to Thm.sort_triv (cf. more_thm.ML);
 wenzelm parents: 
31177diff
changeset | 225 | |
| 31944 | 226 | fun class_triv thy c = | 
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59616diff
changeset | 227 | Thm.of_class (Thm.global_ctyp_of thy (TVar ((Name.aT, 0), [c])), c); | 
| 31944 | 228 | |
| 229 | 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 | 230 | |
| 
a60164e8fff0
added check_shyps, which reject pending sort hypotheses;
 wenzelm parents: 
28116diff
changeset | 231 | |
| 22695 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 wenzelm parents: 
22682diff
changeset | 232 | (* misc operations *) | 
| 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 wenzelm parents: 
22682diff
changeset | 233 | |
| 24048 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
23599diff
changeset | 234 | fun is_dummy thm = | 
| 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
23599diff
changeset | 235 | (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 | 236 | NONE => false | 
| 58001 
934d85f14d1d
more general dummy: may contain "parked arguments", for example;
 wenzelm parents: 
56245diff
changeset | 237 | | 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 | 238 | |
| 22695 
17073e9b94f2
moved Drule.plain_prop_of, Drule.fold_terms to more_thm.ML;
 wenzelm parents: 
22682diff
changeset | 239 | |
| 30564 | 240 | (* collections of theorems in canonical order *) | 
| 24048 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
23599diff
changeset | 241 | |
| 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
23599diff
changeset | 242 | val add_thm = update eq_thm_prop; | 
| 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
23599diff
changeset | 243 | val del_thm = remove eq_thm_prop; | 
| 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
23599diff
changeset | 244 | val merge_thms = merge eq_thm_prop; | 
| 
a12b4faff474
moved Drule.add/del/merge_rules to Thm.add/del/merge_thms;
 wenzelm parents: 
23599diff
changeset | 245 | |
| 74152 | 246 | val item_net = Item_Net.init eq_thm_prop (single o Thm.full_prop_of); | 
| 247 | val item_net_intro = Item_Net.init eq_thm_prop (single o Thm.concl_of); | |
| 248 | val item_net_elim = Item_Net.init eq_thm_prop (single o Thm.major_prem_of); | |
| 30560 | 249 | |
| 250 | ||
| 22682 | 251 | |
| 61508 | 252 | (** declared hyps and sort hyps **) | 
| 54984 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 wenzelm parents: 
53206diff
changeset | 253 | |
| 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 wenzelm parents: 
53206diff
changeset | 254 | structure Hyps = Proof_Data | 
| 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 wenzelm parents: 
53206diff
changeset | 255 | ( | 
| 77869 | 256 |   type T = {checked_hyps: bool, hyps: Termset.T, shyps: sort Ord_List.T};
 | 
| 257 |   fun init _ : T = {checked_hyps = true, hyps = Termset.empty, shyps = []};
 | |
| 54984 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 wenzelm parents: 
53206diff
changeset | 258 | ); | 
| 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 wenzelm parents: 
53206diff
changeset | 259 | |
| 61508 | 260 | fun map_hyps f = Hyps.map (fn {checked_hyps, hyps, shyps} =>
 | 
| 261 | let val (checked_hyps', hyps', shyps') = f (checked_hyps, hyps, shyps) | |
| 262 |   in {checked_hyps = checked_hyps', hyps = hyps', shyps = shyps'} end);
 | |
| 263 | ||
| 264 | ||
| 265 | (* hyps *) | |
| 266 | ||
| 267 | fun declare_hyps raw_ct ctxt = ctxt |> map_hyps (fn (checked_hyps, hyps, shyps) => | |
| 268 | let | |
| 78136 | 269 | val ct = Thm.transfer_cterm' ctxt raw_ct; | 
| 77723 
b761c91c2447
performance tuning: prefer functor Set() over Table();
 wenzelm parents: 
76082diff
changeset | 270 | val hyps' = Termset.insert (Thm.term_of ct) hyps; | 
| 61508 | 271 | in (checked_hyps, hyps', shyps) end); | 
| 54984 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 wenzelm parents: 
53206diff
changeset | 272 | |
| 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 wenzelm parents: 
53206diff
changeset | 273 | 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 | 274 | |
| 61508 | 275 | val unchecked_hyps = map_hyps (fn (_, hyps, shyps) => (false, hyps, shyps)); | 
| 276 | ||
| 277 | fun restore_hyps ctxt = | |
| 278 | 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: 
54984diff
changeset | 279 | |
| 55633 | 280 | fun undeclared_hyps context th = | 
| 77863 
760515c45864
revert b43ee37926a9 due to problems with AFP/PAPP_Impossibility;
 wenzelm parents: 
77808diff
changeset | 281 | Thm.hyps_of th | 
| 
760515c45864
revert b43ee37926a9 due to problems with AFP/PAPP_Impossibility;
 wenzelm parents: 
77808diff
changeset | 282 | |> filter_out | 
| 55633 | 283 | (case context of | 
| 77863 
760515c45864
revert b43ee37926a9 due to problems with AFP/PAPP_Impossibility;
 wenzelm parents: 
77808diff
changeset | 284 | Context.Theory _ => K false | 
| 55633 | 285 | | Context.Proof ctxt => | 
| 77863 
760515c45864
revert b43ee37926a9 due to problems with AFP/PAPP_Impossibility;
 wenzelm parents: 
77808diff
changeset | 286 | (case Hyps.get ctxt of | 
| 
760515c45864
revert b43ee37926a9 due to problems with AFP/PAPP_Impossibility;
 wenzelm parents: 
77808diff
changeset | 287 |           {checked_hyps = false, ...} => K true
 | 
| 
760515c45864
revert b43ee37926a9 due to problems with AFP/PAPP_Impossibility;
 wenzelm parents: 
77808diff
changeset | 288 |         | {hyps, ...} => Termset.member hyps));
 | 
| 55633 | 289 | |
| 54993 
625370769fc0
check_hyps for attribute application (still inactive, due to non-compliant tools);
 wenzelm parents: 
54984diff
changeset | 290 | fun check_hyps context th = | 
| 55633 | 291 | (case undeclared_hyps context th of | 
| 292 | [] => th | |
| 293 | | undeclared => | |
| 61263 | 294 | error (Pretty.string_of (Pretty.big_list "Undeclared hyps:" | 
| 295 | (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: 
53206diff
changeset | 296 | |
| 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 wenzelm parents: 
53206diff
changeset | 297 | |
| 61508 | 298 | (* shyps *) | 
| 299 | ||
| 300 | fun declare_term_sorts t = | |
| 301 | map_hyps (fn (checked_hyps, hyps, shyps) => | |
| 77869 | 302 | (checked_hyps, hyps, Sorts.insert_term t shyps)); | 
| 61508 | 303 | |
| 61509 
358dfae15d83
print thm wrt. local shyps (from full proof context);
 wenzelm parents: 
61508diff
changeset | 304 | fun extra_shyps' ctxt th = | 
| 77869 | 305 | Sorts.subtract (#shyps (Hyps.get ctxt)) (Thm.extra_shyps th); | 
| 61509 
358dfae15d83
print thm wrt. local shyps (from full proof context);
 wenzelm parents: 
61508diff
changeset | 306 | |
| 61508 | 307 | fun check_shyps ctxt raw_th = | 
| 308 | let | |
| 309 | val th = Thm.strip_shyps raw_th; | |
| 61509 
358dfae15d83
print thm wrt. local shyps (from full proof context);
 wenzelm parents: 
61508diff
changeset | 310 | val extra_shyps = extra_shyps' ctxt th; | 
| 61508 | 311 | in | 
| 77869 | 312 | if null extra_shyps then th | 
| 61508 | 313 | else error (Pretty.string_of (Pretty.block (Pretty.str "Pending sort hypotheses:" :: | 
| 77869 | 314 | Pretty.brk 1 :: Pretty.commas (map (Syntax.pretty_sort ctxt) extra_shyps)))) | 
| 61508 | 315 | end; | 
| 316 | ||
| 317 | val weaken_sorts' = Thm.weaken_sorts o #shyps o Hyps.get; | |
| 318 | ||
| 319 | ||
| 54984 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 wenzelm parents: 
53206diff
changeset | 320 | |
| 24980 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 321 | (** basic derived rules **) | 
| 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 322 | |
| 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 323 | (*Elimination of implication | 
| 67721 | 324 | A A \<Longrightarrow> B | 
| 24980 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 325 | ------------ | 
| 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 326 | B | 
| 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 327 | *) | 
| 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 328 | fun elim_implies thA thAB = Thm.implies_elim thAB thA; | 
| 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 329 | |
| 26653 | 330 | |
| 61339 | 331 | (* forall_intr_name *) | 
| 332 | ||
| 333 | fun forall_intr_name (a, x) th = | |
| 334 | let | |
| 335 | val th' = Thm.forall_intr x th; | |
| 336 | val prop' = (case Thm.prop_of th' of all $ Abs (_, T, b) => all $ Abs (a, T, b)); | |
| 337 | in Thm.renamed_prop prop' th' end; | |
| 338 | ||
| 339 | ||
| 26653 | 340 | (* forall_elim_var(s) *) | 
| 341 | ||
| 342 | local | |
| 343 | ||
| 74525 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 344 | val used_frees = | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 345 | Name.build_context o | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 346 |     Thm.fold_terms {hyps = true} Term.declare_term_frees;
 | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 347 | |
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 348 | fun used_vars i = | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 349 | Name.build_context o | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 350 |     (Thm.fold_terms {hyps = false} o Term.fold_aterms)
 | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 351 | (fn Var ((x, j), _) => i = j ? Name.declare x | _ => I); | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 352 | |
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 353 | fun dest_all ct used = | 
| 60951 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 wenzelm parents: 
60950diff
changeset | 354 | (case Thm.term_of ct of | 
| 74525 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 355 |     Const ("Pure.all", _) $ Abs (x, _, _) =>
 | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 356 | let | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 357 | val (x', used') = Name.variant x used; | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 358 | val (v, ct') = Thm.dest_abs_fresh x' (Thm.dest_arg ct); | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 359 | in SOME ((x, Thm.ctyp_of_cterm v), (ct', used')) end | 
| 60951 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 wenzelm parents: 
60950diff
changeset | 360 | | _ => NONE); | 
| 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 wenzelm parents: 
60950diff
changeset | 361 | |
| 74525 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 362 | fun dest_all_list ct used = | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 363 | (case dest_all ct used of | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 364 | NONE => ([], used) | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 365 | | SOME (v, (ct', used')) => | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 366 | let val (vs, used'') = dest_all_list ct' used' | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 367 | in (v :: vs, used'') end); | 
| 60951 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 wenzelm parents: 
60950diff
changeset | 368 | |
| 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 wenzelm parents: 
60950diff
changeset | 369 | fun forall_elim_vars_list vars i th = | 
| 26653 | 370 | let | 
| 74525 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 371 | val (vars', _) = | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 372 | (vars, used_vars i th) |-> fold_map (fn (x, T) => fn used => | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 373 | let val (x', used') = Name.variant x used | 
| 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 374 | in (Thm.var ((x', i), T), used') end); | 
| 60951 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 wenzelm parents: 
60950diff
changeset | 375 | in fold Thm.forall_elim vars' th end; | 
| 26653 | 376 | |
| 377 | in | |
| 378 | ||
| 60950 | 379 | fun forall_elim_vars i th = | 
| 74525 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 380 | forall_elim_vars_list (#1 (dest_all_list (Thm.cprop_of th) (used_frees th))) i th; | 
| 26653 | 381 | |
| 33697 | 382 | fun forall_elim_var i th = | 
| 60950 | 383 | let | 
| 384 | val vars = | |
| 74525 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 385 | (case dest_all (Thm.cprop_of th) (used_frees th) of | 
| 60951 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 wenzelm parents: 
60950diff
changeset | 386 | SOME (v, _) => [v] | 
| 
b70c4db3874f
produce certified vars without access to theory_of_thm, and without context;
 wenzelm parents: 
60950diff
changeset | 387 |       | 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 | 388 | in forall_elim_vars_list vars i th end; | 
| 26653 | 389 | |
| 390 | end; | |
| 391 | ||
| 392 | ||
| 67698 
67caf783b9ee
explicit operations to instantiate frees: typ, term, thm, morphism;
 wenzelm parents: 
67671diff
changeset | 393 | (* instantiate frees *) | 
| 
67caf783b9ee
explicit operations to instantiate frees: typ, term, thm, morphism;
 wenzelm parents: 
67671diff
changeset | 394 | |
| 74282 | 395 | fun instantiate_frees (instT, inst) th = | 
| 396 | if TFrees.is_empty instT andalso Frees.is_empty inst then th | |
| 397 | else | |
| 398 | let | |
| 399 | val idx = Thm.maxidx_of th + 1; | |
| 400 | fun index ((a, A), b) = (((a, idx), A), b); | |
| 401 | val insts = | |
| 402 | (TVars.build (TFrees.fold (TVars.add o index) instT), | |
| 403 | Vars.build (Frees.fold (Vars.add o index) inst)); | |
| 404 | val tfrees = Names.build (TFrees.fold (Names.add_set o #1 o #1) instT); | |
| 405 | val frees = Names.build (Frees.fold (Names.add_set o #1 o #1) inst); | |
| 67698 
67caf783b9ee
explicit operations to instantiate frees: typ, term, thm, morphism;
 wenzelm parents: 
67671diff
changeset | 406 | |
| 74282 | 407 | val hyps = Thm.chyps_of th; | 
| 408 | val inst_cterm = | |
| 409 | Thm.generalize_cterm (tfrees, frees) idx #> | |
| 410 | Thm.instantiate_cterm insts; | |
| 411 | in | |
| 412 | th | |
| 413 | |> fold_rev Thm.implies_intr hyps | |
| 414 | |> Thm.generalize (tfrees, frees) idx | |
| 415 | |> Thm.instantiate insts | |
| 416 | |> fold (elim_implies o Thm.assume o inst_cterm) hyps | |
| 417 | end; | |
| 67698 
67caf783b9ee
explicit operations to instantiate frees: typ, term, thm, morphism;
 wenzelm parents: 
67671diff
changeset | 418 | |
| 
67caf783b9ee
explicit operations to instantiate frees: typ, term, thm, morphism;
 wenzelm parents: 
67671diff
changeset | 419 | |
| 60801 | 420 | (* instantiate by left-to-right occurrence of variables *) | 
| 421 | ||
| 422 | fun instantiate' cTs cts thm = | |
| 423 | let | |
| 424 | fun err msg = | |
| 425 |       raise TYPE ("instantiate': " ^ msg,
 | |
| 426 | map_filter (Option.map Thm.typ_of) cTs, | |
| 427 | map_filter (Option.map Thm.term_of) cts); | |
| 428 | ||
| 429 | fun zip_vars xs ys = | |
| 430 | zip_options xs ys handle ListPair.UnequalLengths => | |
| 431 | err "more instantiations than variables in thm"; | |
| 432 | ||
| 74241 
eb265f54e3ce
more efficient operations: traverse hyps only when required;
 wenzelm parents: 
74239diff
changeset | 433 |     val instT = zip_vars (build_rev (Thm.fold_terms {hyps = false} Term.add_tvars thm)) cTs;
 | 
| 74282 | 434 | val thm' = Thm.instantiate (TVars.make instT, Vars.empty) thm; | 
| 74241 
eb265f54e3ce
more efficient operations: traverse hyps only when required;
 wenzelm parents: 
74239diff
changeset | 435 |     val inst = zip_vars (build_rev (Thm.fold_terms {hyps = false} Term.add_vars thm')) cts;
 | 
| 74282 | 436 | in Thm.instantiate (TVars.empty, Vars.make inst) thm' end; | 
| 60801 | 437 | |
| 438 | ||
| 74245 | 439 | (* implicit generalization over variables -- canonical order *) | 
| 440 | ||
| 441 | fun forall_intr_vars th = | |
| 74271 | 442 | let val vars = Cterms.build (Cterms.add_vars th) | 
| 74269 | 443 | in fold_rev Thm.forall_intr (Cterms.list_set vars) 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 | 444 | |
| 
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 | 445 | 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 | 446 | let | 
| 74246 | 447 | val fixed = | 
| 74266 | 448 | Frees.build | 
| 449 | (fold Frees.add_frees (Thm.terms_of_tpairs (Thm.tpairs_of th)) #> | |
| 77863 
760515c45864
revert b43ee37926a9 due to problems with AFP/PAPP_Impossibility;
 wenzelm parents: 
77808diff
changeset | 450 | fold Frees.add_frees (Thm.hyps_of th)); | 
| 74269 | 451 | val is_fixed = Frees.defined fixed o Term.dest_Free o Thm.term_of; | 
| 452 | val frees = | |
| 453 |       Cterms.build (th |> Thm.fold_atomic_cterms {hyps = false} Term.is_Free
 | |
| 454 | (fn ct => not (is_fixed ct) ? Cterms.add_set ct)); | |
| 455 | in fold_rev Thm.forall_intr (Cterms.list_set 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 | 456 | |
| 
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 | 457 | |
| 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 | 458 | (* unvarify_global: global schematic variables *) | 
| 26653 | 459 | |
| 60825 | 460 | fun unvarify_global thy th = | 
| 24980 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 461 | let | 
| 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 462 | val prop = Thm.full_prop_of th; | 
| 77863 
760515c45864
revert b43ee37926a9 due to problems with AFP/PAPP_Impossibility;
 wenzelm parents: 
77808diff
changeset | 463 | val _ = map Logic.unvarify_global (prop :: Thm.hyps_of th) | 
| 24980 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 464 | handle TERM (msg, _) => raise THM (msg, 0, [th]); | 
| 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 465 | |
| 74230 | 466 | val cert = Thm.global_cterm_of thy; | 
| 467 | val certT = Thm.global_ctyp_of thy; | |
| 468 | ||
| 469 | val instT = | |
| 74266 | 470 | TVars.build (prop |> (Term.fold_types o Term.fold_atyps) | 
| 74230 | 471 | (fn T => fn instT => | 
| 472 | (case T of | |
| 473 | TVar (v as ((a, _), S)) => | |
| 74266 | 474 | if TVars.defined instT v then instT | 
| 475 | else TVars.add (v, TFree (a, S)) instT | |
| 74232 | 476 | | _ => instT))); | 
| 74266 | 477 | val cinstT = TVars.map (K certT) instT; | 
| 74230 | 478 | val cinst = | 
| 74266 | 479 | Vars.build (prop |> Term.fold_aterms | 
| 74230 | 480 | (fn t => fn inst => | 
| 481 | (case t of | |
| 482 | Var ((x, i), T) => | |
| 483 | let val T' = Term_Subst.instantiateT instT T | |
| 74266 | 484 | in Vars.add (((x, i), T'), cert (Free ((x, T')))) inst end | 
| 74232 | 485 | | _ => inst))); | 
| 74282 | 486 | in Thm.instantiate (cinstT, cinst) th end; | 
| 24980 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 487 | |
| 60825 | 488 | fun unvarify_axiom thy = unvarify_global thy o Thm.axiom thy; | 
| 489 | ||
| 26653 | 490 | |
| 59969 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 491 | (* user renaming of parameters in a subgoal *) | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 492 | |
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 493 | (*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 | 494 | preceding parameters may be renamed to make all parameters distinct.*) | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 495 | fun rename_params_rule (names, i) st = | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 496 | let | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 497 | val (_, Bs, Bi, C) = Thm.dest_state (st, i); | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 498 | val params = map #1 (Logic.strip_params Bi); | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 499 | val short = length params - length names; | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 500 | val names' = | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 501 | 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 | 502 | else Name.variant_list names (take short params) @ names; | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 503 | 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 | 504 | val Bi' = Logic.list_rename_params names' Bi; | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 505 | in | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 506 | (case duplicates (op =) names of | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 507 |       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 | 508 | | [] => | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 509 | (case inter (op =) names free_names of | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 510 |         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 | 511 | | [] => Thm.renamed_prop (Logic.list_implies (Bs @ [Bi'], C)) st)) | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 512 | end; | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 513 | |
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 514 | |
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 515 | (* preservation of bound variable names *) | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 516 | |
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 517 | fun rename_boundvars pat obj th = | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 518 | (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 | 519 | NONE => th | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 520 | | SOME prop' => Thm.renamed_prop prop' th); | 
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 521 | |
| 
bcccad156236
explicitly checked alpha conversion -- actual renaming happens outside kernel;
 wenzelm parents: 
59623diff
changeset | 522 | |
| 24980 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 523 | |
| 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 524 | (** specification primitives **) | 
| 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 525 | |
| 30342 | 526 | (* rules *) | 
| 527 | ||
| 35855 
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
 wenzelm parents: 
35853diff
changeset | 528 | fun stripped_sorts thy t = | 
| 
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
 wenzelm parents: 
35853diff
changeset | 529 | let | 
| 74235 | 530 | val tfrees = build_rev (Term.add_tfrees t); | 
| 70922 
539dfd4c166b
more conservative type names, e.g. relevant for Isabelle/MMT export;
 wenzelm parents: 
70915diff
changeset | 531 | val tfrees' = map (fn a => (a, [])) (Name.variant_list [] (map #1 tfrees)); | 
| 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 | 532 | 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 | 533 | 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 | 534 | 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 | 535 | val strip = map (apply2 TFree) (tfrees ~~ tfrees'); | 
| 79436 | 536 | val t' = | 
| 537 | (Term.map_types o Term.map_atyps_same) | |
| 538 | (Same.function_eq (op =) (perhaps (AList.lookup (op =) strip))) t; | |
| 35855 
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
 wenzelm parents: 
35853diff
changeset | 539 | in (strip, recover, t') end; | 
| 
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
 wenzelm parents: 
35853diff
changeset | 540 | |
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
40238diff
changeset | 541 | fun add_axiom ctxt (b, prop) thy = | 
| 24980 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 542 | let | 
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
40238diff
changeset | 543 | val _ = Sign.no_vars ctxt prop; | 
| 35855 
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
 wenzelm parents: 
35853diff
changeset | 544 | val (strip, recover, prop') = stripped_sorts thy prop; | 
| 
e7d004b89ca8
add_axiom: axiomatize "unconstrained" version, with explicit of_class premises;
 wenzelm parents: 
35853diff
changeset | 545 | val constraints = map (fn (TFree (_, S), T) => (T, S)) strip; | 
| 60367 | 546 | 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 | 547 | |
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
40238diff
changeset | 548 | val thy' = thy | 
| 51316 
dfe469293eb4
discontinued empty name bindings in 'axiomatization';
 wenzelm parents: 
49062diff
changeset | 549 | |> 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 | 550 | 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 | 551 | 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 | 552 | val thm = | 
| 74282 | 553 | Thm.instantiate (TVars.make recover, Vars.empty) axm' | 
| 60825 | 554 | |> unvarify_global thy' | 
| 35988 
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
 wenzelm parents: 
35985diff
changeset | 555 | |> fold elim_implies of_sorts; | 
| 36106 
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
 wenzelm parents: 
35988diff
changeset | 556 | in ((axm_name, thm), thy') end; | 
| 24980 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 557 | |
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
40238diff
changeset | 558 | 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 | 559 | |
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61059diff
changeset | 560 | fun add_def (context as (ctxt, _)) unchecked overloaded (b, prop) thy = | 
| 24980 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 561 | let | 
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
40238diff
changeset | 562 | val _ = Sign.no_vars ctxt prop; | 
| 60367 | 563 | 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 | 564 | 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 | 565 | |
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61059diff
changeset | 566 | 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: 
35988diff
changeset | 567 | val axm_name = Sign.full_name thy' b; | 
| 
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
 wenzelm parents: 
35988diff
changeset | 568 | 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 | 569 | val thm = | 
| 74282 | 570 | Thm.instantiate (TVars.make recover, Vars.empty) axm' | 
| 60825 | 571 | |> unvarify_global thy' | 
| 35988 
76ca601c941e
disallow premises in primitive Theory.add_def -- handle in Thm.add_def;
 wenzelm parents: 
35985diff
changeset | 572 | |> fold_rev Thm.implies_intr prems; | 
| 36106 
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
 wenzelm parents: 
35988diff
changeset | 573 | in ((axm_name, thm), thy') end; | 
| 24980 
16a74cfca971
added elim_implies (more convenient argument order);
 wenzelm parents: 
24948diff
changeset | 574 | |
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
40238diff
changeset | 575 | fun add_def_global unchecked overloaded arg thy = | 
| 61262 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61261diff
changeset | 576 | add_def (Defs.global_context thy) unchecked overloaded arg thy; | 
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
40238diff
changeset | 577 | |
| 27866 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 578 | |
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 579 | |
| 70443 | 580 | (** theorem tags **) | 
| 27866 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 581 | |
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 582 | (* add / delete tags *) | 
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 583 | |
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 584 | fun tag_rule tg = Thm.map_tags (insert (op =) tg); | 
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 585 | 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 | 586 | |
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 587 | |
| 61852 | 588 | (* free dummy thm -- for abstract closure *) | 
| 589 | ||
| 590 | val free_dummyN = "free_dummy"; | |
| 591 | fun is_free_dummy thm = Properties.defined (Thm.get_tags thm) free_dummyN; | |
| 592 | val tag_free_dummy = tag_rule (free_dummyN, ""); | |
| 593 | ||
| 594 | ||
| 30342 | 595 | (* def_name *) | 
| 596 | ||
| 597 | fun def_name c = c ^ "_def"; | |
| 598 | ||
| 599 | fun def_name_optional c "" = def_name c | |
| 600 | | def_name_optional _ name = name; | |
| 601 | ||
| 63041 
cb495c4807b3
clarified def binding position: reset for implicit/derived binding, keep for explicit binding;
 wenzelm parents: 
62093diff
changeset | 602 | 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: 
62093diff
changeset | 603 | 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: 
62093diff
changeset | 604 | fun make_def_binding cond b = if cond then def_binding b else Binding.empty; | 
| 62093 | 605 | |
| 30342 | 606 | |
| 27866 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 607 | (* unofficial theorem names *) | 
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 608 | |
| 68036 
4c9e79aeadf0
tuned -- avoid spurious exception trace for "the";
 wenzelm parents: 
67778diff
changeset | 609 | fun has_name_hint thm = AList.defined (op =) (Thm.get_tags thm) Markup.nameN; | 
| 27866 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 610 | fun the_name_hint thm = the (AList.lookup (op =) (Thm.get_tags thm) Markup.nameN); | 
| 68036 
4c9e79aeadf0
tuned -- avoid spurious exception trace for "the";
 wenzelm parents: 
67778diff
changeset | 611 | fun get_name_hint thm = if has_name_hint thm then the_name_hint thm else "??.unknown"; | 
| 27866 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 612 | |
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 613 | 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 | 614 | |
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 615 | |
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 616 | (* theorem kinds *) | 
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 617 | |
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 618 | val theoremK = "theorem"; | 
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 619 | |
| 78462 | 620 | fun legacy_get_kind thm = Properties.get_string (Thm.get_tags thm) Markup.kindN; | 
| 27866 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 621 | |
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 622 | 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: 
61852diff
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: 
61852diff
changeset | 624 | |
| 
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: 
61852diff
changeset | 625 | |
| 
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: 
61852diff
changeset | 626 | (** 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: 
61852diff
changeset | 627 | |
| 
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: 
61852diff
changeset | 628 | (*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: 
61852diff
changeset | 629 | 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: 
61852diff
changeset | 630 | |
| 
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: 
61852diff
changeset | 631 | 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: 
61852diff
changeset | 632 | |
| 
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: 
61852diff
changeset | 633 | 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: 
61852diff
changeset | 634 | (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: 
61852diff
changeset | 635 | (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: 
61852diff
changeset | 636 | 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: 
61852diff
changeset | 637 | | 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: 
61852diff
changeset | 638 | |
| 
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: 
61852diff
changeset | 639 | 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: 
61852diff
changeset | 640 | (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: 
61852diff
changeset | 641 | |
| 
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: 
61852diff
changeset | 642 | 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: 
61852diff
changeset | 643 | 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: 
61852diff
changeset | 644 | |
| 
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: 
61852diff
changeset | 645 | fun apply_attribute (att: attribute) th x = | 
| 67649 | 646 | let val (x', th') = att (x, check_hyps x (Thm.transfer'' x th)) | 
| 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: 
61852diff
changeset | 647 | 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: 
61852diff
changeset | 648 | |
| 
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: 
61852diff
changeset | 649 | 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: 
61852diff
changeset | 650 | |
| 
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: 
61852diff
changeset | 651 | 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: 
61852diff
changeset | 652 | 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: 
61852diff
changeset | 653 | 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: 
61852diff
changeset | 654 | | 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: 
61852diff
changeset | 655 | 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: 
61852diff
changeset | 656 | |
| 
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: 
61852diff
changeset | 657 | 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: 
61852diff
changeset | 658 | 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: 
61852diff
changeset | 659 | |
| 
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: 
61852diff
changeset | 660 | 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: 
61852diff
changeset | 661 | 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: 
61852diff
changeset | 662 | |
| 
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: 
61852diff
changeset | 663 | 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: 
61852diff
changeset | 664 | 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: 
61852diff
changeset | 665 | 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 | 666 | |
| 
c721ea6e0eb4
moved basic thm operations from structure PureThy to Thm;
 wenzelm parents: 
27255diff
changeset | 667 | |
| 70443 | 668 | |
| 669 | (** forked proofs **) | |
| 49010 | 670 | |
| 49062 
7e31dfd99ce7
discontinued complicated/unreliable notion of recent proofs within context;
 wenzelm parents: 
49058diff
changeset | 671 | structure Proofs = Theory_Data | 
| 49010 | 672 | ( | 
| 72048 
d3b8c8b2d1fc
more thorough extend/merge (for Theory.join_theory);
 wenzelm parents: 
71088diff
changeset | 673 | type T = thm list lazy Inttab.table; | 
| 
d3b8c8b2d1fc
more thorough extend/merge (for Theory.join_theory);
 wenzelm parents: 
71088diff
changeset | 674 | val empty = Inttab.empty; | 
| 
d3b8c8b2d1fc
more thorough extend/merge (for Theory.join_theory);
 wenzelm parents: 
71088diff
changeset | 675 | val merge = Inttab.merge (K true); | 
| 49010 | 676 | ); | 
| 677 | ||
| 72048 
d3b8c8b2d1fc
more thorough extend/merge (for Theory.join_theory);
 wenzelm parents: 
71088diff
changeset | 678 | fun reset_proofs thy = | 
| 
d3b8c8b2d1fc
more thorough extend/merge (for Theory.join_theory);
 wenzelm parents: 
71088diff
changeset | 679 | if Inttab.is_empty (Proofs.get thy) then NONE | 
| 
d3b8c8b2d1fc
more thorough extend/merge (for Theory.join_theory);
 wenzelm parents: 
71088diff
changeset | 680 | else SOME (Proofs.put Inttab.empty thy); | 
| 
d3b8c8b2d1fc
more thorough extend/merge (for Theory.join_theory);
 wenzelm parents: 
71088diff
changeset | 681 | |
| 
d3b8c8b2d1fc
more thorough extend/merge (for Theory.join_theory);
 wenzelm parents: 
71088diff
changeset | 682 | val _ = Theory.setup (Theory.at_begin reset_proofs); | 
| 61059 | 683 | |
| 72048 
d3b8c8b2d1fc
more thorough extend/merge (for Theory.join_theory);
 wenzelm parents: 
71088diff
changeset | 684 | fun register_proofs ths thy = | 
| 
d3b8c8b2d1fc
more thorough extend/merge (for Theory.join_theory);
 wenzelm parents: 
71088diff
changeset | 685 | let val entry = (serial (), Lazy.map_finished (map Thm.trim_context) ths) | 
| 
d3b8c8b2d1fc
more thorough extend/merge (for Theory.join_theory);
 wenzelm parents: 
71088diff
changeset | 686 | in (Proofs.map o Inttab.update) entry thy end; | 
| 
d3b8c8b2d1fc
more thorough extend/merge (for Theory.join_theory);
 wenzelm parents: 
71088diff
changeset | 687 | |
| 
d3b8c8b2d1fc
more thorough extend/merge (for Theory.join_theory);
 wenzelm parents: 
71088diff
changeset | 688 | fun force_proofs thy = | 
| 
d3b8c8b2d1fc
more thorough extend/merge (for Theory.join_theory);
 wenzelm parents: 
71088diff
changeset | 689 | Proofs.get thy |> Inttab.dest |> maps (map (Thm.transfer thy) o Lazy.force o #2); | 
| 71023 
35a8e15b7e03
more robust Thm.expose_theory -- ensure that PIDE export happens in the proper theory context;
 wenzelm parents: 
71018diff
changeset | 690 | |
| 
35a8e15b7e03
more robust Thm.expose_theory -- ensure that PIDE export happens in the proper theory context;
 wenzelm parents: 
71018diff
changeset | 691 | val consolidate_theory = Thm.consolidate o force_proofs; | 
| 
35a8e15b7e03
more robust Thm.expose_theory -- ensure that PIDE export happens in the proper theory context;
 wenzelm parents: 
71018diff
changeset | 692 | |
| 
35a8e15b7e03
more robust Thm.expose_theory -- ensure that PIDE export happens in the proper theory context;
 wenzelm parents: 
71018diff
changeset | 693 | fun expose_theory thy = | 
| 
35a8e15b7e03
more robust Thm.expose_theory -- ensure that PIDE export happens in the proper theory context;
 wenzelm parents: 
71018diff
changeset | 694 | if Proofterm.export_enabled () | 
| 
35a8e15b7e03
more robust Thm.expose_theory -- ensure that PIDE export happens in the proper theory context;
 wenzelm parents: 
71018diff
changeset | 695 | then Thm.expose_proofs thy (force_proofs thy) else (); | 
| 49010 | 696 | |
| 697 | ||
| 61268 | 698 | |
| 699 | (** print theorems **) | |
| 700 | ||
| 701 | (* options *) | |
| 702 | ||
| 69575 | 703 | val show_consts = Config.declare_option_bool ("show_consts", \<^here>);
 | 
| 704 | val show_hyps = Config.declare_bool ("show_hyps", \<^here>) (K false);
 | |
| 705 | val show_tags = Config.declare_bool ("show_tags", \<^here>) (K false);
 | |
| 61268 | 706 | |
| 707 | ||
| 708 | (* pretty_thm etc. *) | |
| 709 | ||
| 710 | fun pretty_tag (name, arg) = Pretty.strs [name, quote arg]; | |
| 711 | val pretty_tags = Pretty.list "[" "]" o map pretty_tag; | |
| 712 | ||
| 713 | fun pretty_thm_raw ctxt {quote, show_hyps = show_hyps'} raw_th =
 | |
| 714 | let | |
| 715 | val show_tags = Config.get ctxt show_tags; | |
| 716 | val show_hyps = Config.get ctxt show_hyps; | |
| 717 | ||
| 718 | val th = raw_th | |
| 67649 | 719 | |> perhaps (try (Thm.transfer' ctxt)) | 
| 61268 | 720 | |> perhaps (try Thm.strip_shyps); | 
| 721 | ||
| 77863 
760515c45864
revert b43ee37926a9 due to problems with AFP/PAPP_Impossibility;
 wenzelm parents: 
77808diff
changeset | 722 | val hyps = if show_hyps then Thm.hyps_of th else undeclared_hyps (Context.Proof ctxt) th; | 
| 77869 | 723 | val extra_shyps = extra_shyps' ctxt th; | 
| 61268 | 724 | val tags = Thm.get_tags th; | 
| 725 | val tpairs = Thm.tpairs_of th; | |
| 726 | ||
| 727 | val q = if quote then Pretty.quote else I; | |
| 728 | val prt_term = q o Syntax.pretty_term ctxt; | |
| 729 | ||
| 730 | ||
| 731 | val hlen = length extra_shyps + length hyps + length tpairs; | |
| 732 | val hsymbs = | |
| 733 | if hlen = 0 then [] | |
| 734 | else if show_hyps orelse show_hyps' then | |
| 735 | [Pretty.brk 2, Pretty.list "[" "]" | |
| 76082 | 736 | (map (q o Pretty.block o Syntax.pretty_flexpair ctxt) tpairs @ map prt_term hyps @ | 
| 61268 | 737 | map (Syntax.pretty_sort ctxt) extra_shyps)] | 
| 738 |       else [Pretty.brk 2, Pretty.str ("[" ^ replicate_string hlen "." ^ "]")];
 | |
| 739 | val tsymbs = | |
| 740 | if null tags orelse not show_tags then [] | |
| 741 | else [Pretty.brk 1, pretty_tags tags]; | |
| 742 | in Pretty.block (prt_term (Thm.prop_of th) :: (hsymbs @ tsymbs)) end; | |
| 743 | ||
| 744 | fun pretty_thm ctxt = pretty_thm_raw ctxt {quote = false, show_hyps = true};
 | |
| 745 | fun pretty_thm_item ctxt th = Pretty.item [pretty_thm ctxt th]; | |
| 746 | ||
| 747 | fun pretty_thm_global thy = | |
| 748 |   pretty_thm_raw (Syntax.init_pretty_global thy) {quote = false, show_hyps = false};
 | |
| 749 | ||
| 750 | val string_of_thm = Pretty.string_of oo pretty_thm; | |
| 751 | val string_of_thm_global = Pretty.string_of oo pretty_thm_global; | |
| 752 | ||
| 753 | ||
| 22362 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 754 | open Thm; | 
| 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 755 | |
| 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 756 | end; | 
| 
6470ce514b6e
Further operations on type thm, outside the inference kernel.
 wenzelm parents: diff
changeset | 757 | |
| 32842 | 758 | structure Basic_Thm: BASIC_THM = Thm; | 
| 759 | open Basic_Thm; |