author | wenzelm |
Sun, 16 Jun 2024 21:53:24 +0200 | |
changeset 80395 | 46135b44b1a3 |
parent 78138 | 0ea55458f867 |
permissions | -rw-r--r-- |
20222 | 1 |
(* Title: Pure/assumption.ML |
2 |
Author: Makarius |
|
3 |
||
30471
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
4 |
Context assumptions, parameterized by export rules. |
20222 | 5 |
*) |
6 |
||
7 |
signature ASSUMPTION = |
|
8 |
sig |
|
35716 | 9 |
type export = bool -> cterm list -> (thm -> thm) * (term -> term) |
20222 | 10 |
val assume_export: export |
11 |
val presume_export: export |
|
54883
dd04a8b654fc
proper context for norm_hhf and derived operations;
wenzelm
parents:
54740
diff
changeset
|
12 |
val assume: Proof.context -> cterm -> thm |
54984
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
wenzelm
parents:
54883
diff
changeset
|
13 |
val assume_hyps: cterm -> Proof.context -> thm * Proof.context |
30471
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
14 |
val all_assms_of: Proof.context -> cterm list |
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
15 |
val all_prems_of: Proof.context -> thm list |
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
16 |
val local_assms_of: Proof.context -> Proof.context -> cterm list |
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
17 |
val local_prems_of: Proof.context -> Proof.context -> thm list |
20296 | 18 |
val add_assms: export -> cterm list -> Proof.context -> thm list * Proof.context |
19 |
val add_assumes: cterm list -> Proof.context -> thm list * Proof.context |
|
78084
f0aca0506531
more robust context: fail immediately via Morphism.the_theory, instead of rarely via Thm.theory_of_thm (for non-normal thm);
wenzelm
parents:
78062
diff
changeset
|
20 |
val export_term: Proof.context -> Proof.context -> term -> term |
78086 | 21 |
val export_: {goal: bool} -> Proof.context -> Proof.context -> thm -> thm |
22 |
val export: Proof.context -> Proof.context -> thm -> thm |
|
23 |
val export_goal: Proof.context -> Proof.context -> thm -> thm |
|
21517 | 24 |
val export_morphism: Proof.context -> Proof.context -> morphism |
20222 | 25 |
end; |
26 |
||
27 |
structure Assumption: ASSUMPTION = |
|
28 |
struct |
|
29 |
||
78137 | 30 |
(** export operations **) |
20222 | 31 |
|
21679 | 32 |
type export = bool -> cterm list -> (thm -> thm) * (term -> term); |
20222 | 33 |
|
78137 | 34 |
val term_export = fold_rev (fn (e: export, As) => #2 (e false As)); |
35 |
fun thm_export is_goal = fold_rev (fn (e: export, As) => #1 (e is_goal As)); |
|
36 |
||
37 |
||
38 |
||
39 |
(** basic rules **) |
|
40 |
||
20222 | 41 |
(* |
42 |
[A] |
|
43 |
: |
|
44 |
B |
|
45 |
-------- |
|
67721 | 46 |
#A \<Longrightarrow> B |
20222 | 47 |
*) |
21679 | 48 |
fun assume_export is_goal asms = |
49 |
(if is_goal then Drule.implies_intr_protected asms else Drule.implies_intr_list asms, fn t => t); |
|
20222 | 50 |
|
51 |
(* |
|
52 |
[A] |
|
53 |
: |
|
54 |
B |
|
55 |
------- |
|
67721 | 56 |
A \<Longrightarrow> B |
20222 | 57 |
*) |
58 |
fun presume_export _ = assume_export false; |
|
59 |
||
54984
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
wenzelm
parents:
54883
diff
changeset
|
60 |
|
54883
dd04a8b654fc
proper context for norm_hhf and derived operations;
wenzelm
parents:
54740
diff
changeset
|
61 |
fun assume ctxt = Raw_Simplifier.norm_hhf ctxt o Thm.assume; |
20222 | 62 |
|
54984
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
wenzelm
parents:
54883
diff
changeset
|
63 |
fun assume_hyps ct ctxt = |
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
wenzelm
parents:
54883
diff
changeset
|
64 |
let val (th, ctxt') = Thm.assume_hyps ct ctxt |
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
wenzelm
parents:
54883
diff
changeset
|
65 |
in (Raw_Simplifier.norm_hhf ctxt' th, ctxt') end; |
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
wenzelm
parents:
54883
diff
changeset
|
66 |
|
20222 | 67 |
|
68 |
||
69 |
(** local context data **) |
|
70 |
||
71 |
datatype data = Data of |
|
78134
a11ebc8c751a
minor performance tuning: avoid append to end-of-list;
wenzelm
parents:
78086
diff
changeset
|
72 |
{rev_assms: (export * cterm list) list, (*assumes: A \<Longrightarrow> _*) |
a11ebc8c751a
minor performance tuning: avoid append to end-of-list;
wenzelm
parents:
78086
diff
changeset
|
73 |
rev_prems: thm list}; (*prems: A |- norm_hhf A*) |
20222 | 74 |
|
78134
a11ebc8c751a
minor performance tuning: avoid append to end-of-list;
wenzelm
parents:
78086
diff
changeset
|
75 |
fun make_data (rev_assms, rev_prems) = Data {rev_assms = rev_assms, rev_prems = rev_prems}; |
59150 | 76 |
val empty_data = make_data ([], []); |
20222 | 77 |
|
33519 | 78 |
structure Data = Proof_Data |
20222 | 79 |
( |
80 |
type T = data; |
|
59150 | 81 |
fun init _ = empty_data; |
20222 | 82 |
); |
83 |
||
78134
a11ebc8c751a
minor performance tuning: avoid append to end-of-list;
wenzelm
parents:
78086
diff
changeset
|
84 |
fun map_data f = Data.map (fn Data {rev_assms, rev_prems} => make_data (f (rev_assms, rev_prems))); |
20222 | 85 |
|
30471
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
86 |
|
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
87 |
(* all assumptions *) |
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
88 |
|
78138
0ea55458f867
proper trim_context / transfer, e.g. for Specification.definition;
wenzelm
parents:
78137
diff
changeset
|
89 |
fun all_assumptions_of ctxt = |
0ea55458f867
proper trim_context / transfer, e.g. for Specification.definition;
wenzelm
parents:
78137
diff
changeset
|
90 |
let val Data {rev_assms, ...} = Data.get ctxt |
0ea55458f867
proper trim_context / transfer, e.g. for Specification.definition;
wenzelm
parents:
78137
diff
changeset
|
91 |
in fold (cons o (apsnd o map) (Thm.transfer_cterm' ctxt)) rev_assms [] end; |
0ea55458f867
proper trim_context / transfer, e.g. for Specification.definition;
wenzelm
parents:
78137
diff
changeset
|
92 |
|
0ea55458f867
proper trim_context / transfer, e.g. for Specification.definition;
wenzelm
parents:
78137
diff
changeset
|
93 |
fun all_prems_of ctxt = |
0ea55458f867
proper trim_context / transfer, e.g. for Specification.definition;
wenzelm
parents:
78137
diff
changeset
|
94 |
let val Data {rev_prems, ...} = Data.get ctxt |
0ea55458f867
proper trim_context / transfer, e.g. for Specification.definition;
wenzelm
parents:
78137
diff
changeset
|
95 |
in fold (cons o Thm.transfer' ctxt) rev_prems [] end; |
0ea55458f867
proper trim_context / transfer, e.g. for Specification.definition;
wenzelm
parents:
78137
diff
changeset
|
96 |
|
30471
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
97 |
val all_assms_of = maps #2 o all_assumptions_of; |
20222 | 98 |
|
99 |
||
30471
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
100 |
(* local assumptions *) |
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
101 |
|
70735 | 102 |
local |
103 |
||
104 |
fun drop_prefix eq (args as (x :: xs, y :: ys)) = |
|
105 |
if eq (x, y) then drop_prefix eq (xs, ys) else args |
|
106 |
| drop_prefix _ args = args; |
|
107 |
||
108 |
fun check_result ctxt kind term_of res = |
|
109 |
(case res of |
|
110 |
([], rest) => rest |
|
111 |
| (bad :: _, _) => |
|
112 |
raise Fail ("Outer context disagrees on " ^ kind ^ ": " ^ |
|
113 |
Syntax.string_of_term ctxt (term_of bad))); |
|
114 |
||
115 |
in |
|
116 |
||
30471
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
117 |
fun local_assumptions_of inner outer = |
70735 | 118 |
drop_prefix (eq_snd (eq_list Thm.aconvc)) (apply2 all_assumptions_of (outer, inner)) |
119 |
|>> maps #2 |
|
120 |
|> check_result outer "assumption" Thm.term_of; |
|
26392 | 121 |
|
30471
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
122 |
val local_assms_of = maps #2 oo local_assumptions_of; |
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
123 |
|
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
124 |
fun local_prems_of inner outer = |
70735 | 125 |
drop_prefix Thm.eq_thm_prop (apply2 all_prems_of (outer, inner)) |
126 |
|> check_result outer "premise" Thm.prop_of; |
|
127 |
||
128 |
end; |
|
26392 | 129 |
|
130 |
||
20222 | 131 |
(* add assumptions *) |
132 |
||
54883
dd04a8b654fc
proper context for norm_hhf and derived operations;
wenzelm
parents:
54740
diff
changeset
|
133 |
fun add_assms export new_asms ctxt = |
54984
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
wenzelm
parents:
54883
diff
changeset
|
134 |
let val (new_prems, ctxt') = fold_map assume_hyps new_asms ctxt in |
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
wenzelm
parents:
54883
diff
changeset
|
135 |
ctxt' |
78134
a11ebc8c751a
minor performance tuning: avoid append to end-of-list;
wenzelm
parents:
78086
diff
changeset
|
136 |
|> map_data (fn (rev_assms, rev_prems) => |
78138
0ea55458f867
proper trim_context / transfer, e.g. for Specification.definition;
wenzelm
parents:
78137
diff
changeset
|
137 |
((export, map Thm.trim_context_cterm new_asms) :: rev_assms, |
0ea55458f867
proper trim_context / transfer, e.g. for Specification.definition;
wenzelm
parents:
78137
diff
changeset
|
138 |
fold (cons o Thm.trim_context) new_prems rev_prems)) |
54883
dd04a8b654fc
proper context for norm_hhf and derived operations;
wenzelm
parents:
54740
diff
changeset
|
139 |
|> pair new_prems |
20222 | 140 |
end; |
141 |
||
142 |
val add_assumes = add_assms assume_export; |
|
143 |
||
144 |
||
20296 | 145 |
(* export *) |
20222 | 146 |
|
21679 | 147 |
fun export_term inner outer = |
78137 | 148 |
term_export (local_assumptions_of inner outer); |
21679 | 149 |
|
78084
f0aca0506531
more robust context: fail immediately via Morphism.the_theory, instead of rarely via Thm.theory_of_thm (for non-normal thm);
wenzelm
parents:
78062
diff
changeset
|
150 |
fun export_thm is_goal inner outer = |
78137 | 151 |
thm_export is_goal (local_assumptions_of inner outer); |
78084
f0aca0506531
more robust context: fail immediately via Morphism.the_theory, instead of rarely via Thm.theory_of_thm (for non-normal thm);
wenzelm
parents:
78062
diff
changeset
|
152 |
|
78086 | 153 |
fun export_{goal} inner outer = |
78084
f0aca0506531
more robust context: fail immediately via Morphism.the_theory, instead of rarely via Thm.theory_of_thm (for non-normal thm);
wenzelm
parents:
78062
diff
changeset
|
154 |
Raw_Simplifier.norm_hhf_protect inner #> |
78086 | 155 |
export_thm goal inner outer #> |
78084
f0aca0506531
more robust context: fail immediately via Morphism.the_theory, instead of rarely via Thm.theory_of_thm (for non-normal thm);
wenzelm
parents:
78062
diff
changeset
|
156 |
Raw_Simplifier.norm_hhf_protect outer; |
f0aca0506531
more robust context: fail immediately via Morphism.the_theory, instead of rarely via Thm.theory_of_thm (for non-normal thm);
wenzelm
parents:
78062
diff
changeset
|
157 |
|
78086 | 158 |
val export = export_{goal = false}; |
159 |
val export_goal = export_{goal = true}; |
|
160 |
||
21517 | 161 |
fun export_morphism inner outer = |
162 |
let |
|
78138
0ea55458f867
proper trim_context / transfer, e.g. for Specification.definition;
wenzelm
parents:
78137
diff
changeset
|
163 |
val assms0 = local_assumptions_of inner outer |
0ea55458f867
proper trim_context / transfer, e.g. for Specification.definition;
wenzelm
parents:
78137
diff
changeset
|
164 |
|> (map o apsnd o map) Thm.trim_context_cterm; |
78084
f0aca0506531
more robust context: fail immediately via Morphism.the_theory, instead of rarely via Thm.theory_of_thm (for non-normal thm);
wenzelm
parents:
78062
diff
changeset
|
165 |
fun thm thy = |
f0aca0506531
more robust context: fail immediately via Morphism.the_theory, instead of rarely via Thm.theory_of_thm (for non-normal thm);
wenzelm
parents:
78062
diff
changeset
|
166 |
let val norm = norm_hhf_protect (Proof_Context.init_global thy) |
78138
0ea55458f867
proper trim_context / transfer, e.g. for Specification.definition;
wenzelm
parents:
78137
diff
changeset
|
167 |
in norm #> thm_export false assms0 #> norm end; |
0ea55458f867
proper trim_context / transfer, e.g. for Specification.definition;
wenzelm
parents:
78137
diff
changeset
|
168 |
val term = term_export assms0; |
21517 | 169 |
val typ = Logic.type_map term; |
54740 | 170 |
in |
171 |
Morphism.morphism "Assumption.export" |
|
78084
f0aca0506531
more robust context: fail immediately via Morphism.the_theory, instead of rarely via Thm.theory_of_thm (for non-normal thm);
wenzelm
parents:
78062
diff
changeset
|
172 |
{binding = [], typ = [K typ], term = [K term], fact = [map o thm o Morphism.the_theory]} |
f0aca0506531
more robust context: fail immediately via Morphism.the_theory, instead of rarely via Thm.theory_of_thm (for non-normal thm);
wenzelm
parents:
78062
diff
changeset
|
173 |
|> Morphism.set_context (Proof_Context.theory_of inner) |
54740 | 174 |
end; |
21517 | 175 |
|
20222 | 176 |
end; |