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