author | wenzelm |
Mon, 11 Nov 2013 17:34:44 +0100 | |
changeset 54383 | 9d3c7a04a65e |
parent 47236 | 973ab740a25d |
child 54567 | cfe53047dc16 |
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 |
|
12 |
val assume: cterm -> thm |
|
30471
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
13 |
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
|
14 |
val all_prems_of: Proof.context -> thm list |
20296 | 15 |
val extra_hyps: Proof.context -> thm -> term list |
30471
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 |
||
41228
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents:
39557
diff
changeset
|
51 |
val assume = Raw_Simplifier.norm_hhf o Thm.assume; |
20222 | 52 |
|
53 |
||
54 |
||
55 |
(** local context data **) |
|
56 |
||
57 |
datatype data = Data of |
|
30479 | 58 |
{assms: (export * cterm list) list, (*assumes: A ==> _*) |
47236 | 59 |
prems: thm list}; (*prems: A |- norm_hhf A*) |
20222 | 60 |
|
61 |
fun make_data (assms, prems) = Data {assms = assms, prems = prems}; |
|
62 |
||
33519 | 63 |
structure Data = Proof_Data |
20222 | 64 |
( |
65 |
type T = data; |
|
66 |
fun init _ = make_data ([], []); |
|
67 |
); |
|
68 |
||
69 |
fun map_data f = Data.map (fn Data {assms, prems} => make_data (f (assms, prems))); |
|
45650 | 70 |
fun rep_data ctxt = Data.get ctxt |> (fn Data rep => rep); |
20222 | 71 |
|
30471
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
72 |
|
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
73 |
(* all assumptions *) |
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
74 |
|
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
75 |
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
|
76 |
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
|
77 |
val all_prems_of = #prems o rep_data; |
20222 | 78 |
|
30471
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
79 |
fun extra_hyps ctxt th = |
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
80 |
subtract (op aconv) (map Thm.term_of (all_assms_of ctxt)) (Thm.hyps_of th); |
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
81 |
|
20222 | 82 |
|
30471
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
83 |
(* local assumptions *) |
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
84 |
|
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
85 |
fun local_assumptions_of inner outer = |
33957 | 86 |
drop (length (all_assumptions_of outer)) (all_assumptions_of inner); |
26392 | 87 |
|
30471
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
88 |
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
|
89 |
|
178de3995e91
renamed assms_of to all_assms_of, and prems_of to all_prems_of;
wenzelm
parents:
29605
diff
changeset
|
90 |
fun local_prems_of inner outer = |
33957 | 91 |
drop (length (all_prems_of outer)) (all_prems_of inner); |
26392 | 92 |
|
93 |
||
20222 | 94 |
(* add assumptions *) |
95 |
||
96 |
fun add_assms export new_asms = |
|
97 |
let val new_prems = map assume new_asms in |
|
98 |
map_data (fn (asms, prems) => (asms @ [(export, new_asms)], prems @ new_prems)) #> |
|
99 |
pair new_prems |
|
100 |
end; |
|
101 |
||
102 |
val add_assumes = add_assms assume_export; |
|
103 |
||
104 |
||
20296 | 105 |
(* export *) |
20222 | 106 |
|
20296 | 107 |
fun export is_goal inner outer = |
41228
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents:
39557
diff
changeset
|
108 |
Raw_Simplifier.norm_hhf_protect #> |
30479 | 109 |
fold_rev (fn (e, As) => #1 (e is_goal As)) (local_assumptions_of inner outer) #> |
41228
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents:
39557
diff
changeset
|
110 |
Raw_Simplifier.norm_hhf_protect; |
20222 | 111 |
|
21679 | 112 |
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
|
113 |
fold_rev (fn (e, As) => #2 (e false As)) (local_assumptions_of inner outer); |
21679 | 114 |
|
21517 | 115 |
fun export_morphism inner outer = |
116 |
let |
|
117 |
val thm = export false inner outer; |
|
21679 | 118 |
val term = export_term inner outer; |
21517 | 119 |
val typ = Logic.type_map term; |
45289
25e9e7f527b4
slightly more explicit/syntactic modelling of morphisms;
wenzelm
parents:
41552
diff
changeset
|
120 |
in Morphism.morphism {binding = [], typ = [typ], term = [term], fact = [map thm]} end; |
21517 | 121 |
|
20222 | 122 |
end; |