author | wenzelm |
Sun, 18 Jan 2015 12:50:36 +0100 | |
changeset 59389 | c427f3de9050 |
parent 59150 | 71b416020f42 |
child 67721 | 5348bea4accd |
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 |
|
20 |
val export: bool -> Proof.context -> Proof.context -> thm -> thm |
|
21679 | 21 |
val export_term: Proof.context -> Proof.context -> term -> term |
21517 | 22 |
val export_morphism: Proof.context -> Proof.context -> morphism |
20222 | 23 |
end; |
24 |
||
25 |
structure Assumption: ASSUMPTION = |
|
26 |
struct |
|
27 |
||
28 |
(** basic rules **) |
|
29 |
||
21679 | 30 |
type export = bool -> cterm list -> (thm -> thm) * (term -> term); |
20222 | 31 |
|
32 |
(* |
|
33 |
[A] |
|
34 |
: |
|
35 |
B |
|
36 |
-------- |
|
37 |
#A ==> B |
|
38 |
*) |
|
21679 | 39 |
fun assume_export is_goal asms = |
40 |
(if is_goal then Drule.implies_intr_protected asms else Drule.implies_intr_list asms, fn t => t); |
|
20222 | 41 |
|
42 |
(* |
|
43 |
[A] |
|
44 |
: |
|
45 |
B |
|
46 |
------- |
|
47 |
A ==> B |
|
48 |
*) |
|
49 |
fun presume_export _ = assume_export false; |
|
50 |
||
54984
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
wenzelm
parents:
54883
diff
changeset
|
51 |
|
54883
dd04a8b654fc
proper context for norm_hhf and derived operations;
wenzelm
parents:
54740
diff
changeset
|
52 |
fun assume ctxt = Raw_Simplifier.norm_hhf ctxt o Thm.assume; |
20222 | 53 |
|
54984
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
wenzelm
parents:
54883
diff
changeset
|
54 |
fun assume_hyps ct ctxt = |
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
wenzelm
parents:
54883
diff
changeset
|
55 |
let val (th, ctxt') = Thm.assume_hyps ct ctxt |
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
wenzelm
parents:
54883
diff
changeset
|
56 |
in (Raw_Simplifier.norm_hhf ctxt' th, ctxt') end; |
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
wenzelm
parents:
54883
diff
changeset
|
57 |
|
20222 | 58 |
|
59 |
||
60 |
(** local context data **) |
|
61 |
||
62 |
datatype data = Data of |
|
30479 | 63 |
{assms: (export * cterm list) list, (*assumes: A ==> _*) |
47236 | 64 |
prems: thm list}; (*prems: A |- norm_hhf A*) |
20222 | 65 |
|
66 |
fun make_data (assms, prems) = Data {assms = assms, prems = prems}; |
|
59150 | 67 |
val empty_data = make_data ([], []); |
20222 | 68 |
|
33519 | 69 |
structure Data = Proof_Data |
20222 | 70 |
( |
71 |
type T = data; |
|
59150 | 72 |
fun init _ = empty_data; |
20222 | 73 |
); |
74 |
||
75 |
fun map_data f = Data.map (fn Data {assms, prems} => make_data (f (assms, prems))); |
|
45650 | 76 |
fun rep_data ctxt = Data.get ctxt |> (fn Data rep => rep); |
20222 | 77 |
|
30471
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
78 |
|
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
79 |
(* all assumptions *) |
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
80 |
|
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
81 |
val all_assumptions_of = #assms o rep_data; |
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
82 |
val all_assms_of = maps #2 o all_assumptions_of; |
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
83 |
val all_prems_of = #prems o rep_data; |
20222 | 84 |
|
85 |
||
30471
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
86 |
(* local assumptions *) |
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
87 |
|
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
88 |
fun local_assumptions_of inner outer = |
33957 | 89 |
drop (length (all_assumptions_of outer)) (all_assumptions_of inner); |
26392 | 90 |
|
30471
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
91 |
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
|
92 |
|
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
93 |
fun local_prems_of inner outer = |
33957 | 94 |
drop (length (all_prems_of outer)) (all_prems_of inner); |
26392 | 95 |
|
96 |
||
20222 | 97 |
(* add assumptions *) |
98 |
||
54883
dd04a8b654fc
proper context for norm_hhf and derived operations;
wenzelm
parents:
54740
diff
changeset
|
99 |
fun add_assms export new_asms ctxt = |
54984
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
wenzelm
parents:
54883
diff
changeset
|
100 |
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
|
101 |
ctxt' |
54883
dd04a8b654fc
proper context for norm_hhf and derived operations;
wenzelm
parents:
54740
diff
changeset
|
102 |
|> map_data (fn (asms, prems) => (asms @ [(export, new_asms)], prems @ new_prems)) |
dd04a8b654fc
proper context for norm_hhf and derived operations;
wenzelm
parents:
54740
diff
changeset
|
103 |
|> pair new_prems |
20222 | 104 |
end; |
105 |
||
106 |
val add_assumes = add_assms assume_export; |
|
107 |
||
108 |
||
20296 | 109 |
(* export *) |
20222 | 110 |
|
20296 | 111 |
fun export is_goal inner outer = |
54883
dd04a8b654fc
proper context for norm_hhf and derived operations;
wenzelm
parents:
54740
diff
changeset
|
112 |
Raw_Simplifier.norm_hhf_protect inner #> |
30479 | 113 |
fold_rev (fn (e, As) => #1 (e is_goal As)) (local_assumptions_of inner outer) #> |
54883
dd04a8b654fc
proper context for norm_hhf and derived operations;
wenzelm
parents:
54740
diff
changeset
|
114 |
Raw_Simplifier.norm_hhf_protect outer; |
20222 | 115 |
|
21679 | 116 |
fun export_term inner outer = |
30471
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
117 |
fold_rev (fn (e, As) => #2 (e false As)) (local_assumptions_of inner outer); |
21679 | 118 |
|
21517 | 119 |
fun export_morphism inner outer = |
120 |
let |
|
121 |
val thm = export false inner outer; |
|
21679 | 122 |
val term = export_term inner outer; |
21517 | 123 |
val typ = Logic.type_map term; |
54740 | 124 |
in |
125 |
Morphism.morphism "Assumption.export" |
|
126 |
{binding = [], typ = [typ], term = [term], fact = [map thm]} |
|
127 |
end; |
|
21517 | 128 |
|
20222 | 129 |
end; |