author | wenzelm |
Mon, 15 Jan 2018 14:31:57 +0100 | |
changeset 67438 | fdb7b995974d |
parent 61060 | a2c6f7f64aca |
child 67649 | 1e1782c1aedf |
permissions | -rw-r--r-- |
33374
8099185908a4
Rules that characterize functional/relational specifications.
wenzelm
parents:
diff
changeset
|
1 |
(* Title: Pure/Isar/spec_rules.ML |
8099185908a4
Rules that characterize functional/relational specifications.
wenzelm
parents:
diff
changeset
|
2 |
Author: Makarius |
8099185908a4
Rules that characterize functional/relational specifications.
wenzelm
parents:
diff
changeset
|
3 |
|
33454 | 4 |
Rules that characterize specifications, with rough classification. |
5 |
NB: In the face of arbitrary morphisms, the original shape of |
|
6 |
specifications may get lost. |
|
33374
8099185908a4
Rules that characterize functional/relational specifications.
wenzelm
parents:
diff
changeset
|
7 |
*) |
8099185908a4
Rules that characterize functional/relational specifications.
wenzelm
parents:
diff
changeset
|
8 |
|
8099185908a4
Rules that characterize functional/relational specifications.
wenzelm
parents:
diff
changeset
|
9 |
signature SPEC_RULES = |
8099185908a4
Rules that characterize functional/relational specifications.
wenzelm
parents:
diff
changeset
|
10 |
sig |
33454 | 11 |
datatype rough_classification = Unknown | Equational | Inductive | Co_Inductive |
12 |
type spec = rough_classification * (term list * thm list) |
|
13 |
val get: Proof.context -> spec list |
|
14 |
val get_global: theory -> spec list |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33702
diff
changeset
|
15 |
val retrieve: Proof.context -> term -> spec list |
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33702
diff
changeset
|
16 |
val retrieve_global: theory -> term -> spec list |
33454 | 17 |
val add: rough_classification -> term list * thm list -> local_theory -> local_theory |
18 |
val add_global: rough_classification -> term list * thm list -> theory -> theory |
|
33374
8099185908a4
Rules that characterize functional/relational specifications.
wenzelm
parents:
diff
changeset
|
19 |
end; |
8099185908a4
Rules that characterize functional/relational specifications.
wenzelm
parents:
diff
changeset
|
20 |
|
8099185908a4
Rules that characterize functional/relational specifications.
wenzelm
parents:
diff
changeset
|
21 |
structure Spec_Rules: SPEC_RULES = |
8099185908a4
Rules that characterize functional/relational specifications.
wenzelm
parents:
diff
changeset
|
22 |
struct |
8099185908a4
Rules that characterize functional/relational specifications.
wenzelm
parents:
diff
changeset
|
23 |
|
61060 | 24 |
(* rules *) |
33374
8099185908a4
Rules that characterize functional/relational specifications.
wenzelm
parents:
diff
changeset
|
25 |
|
33454 | 26 |
datatype rough_classification = Unknown | Equational | Inductive | Co_Inductive; |
27 |
type spec = rough_classification * (term list * thm list); |
|
33374
8099185908a4
Rules that characterize functional/relational specifications.
wenzelm
parents:
diff
changeset
|
28 |
|
33519 | 29 |
structure Rules = Generic_Data |
33374
8099185908a4
Rules that characterize functional/relational specifications.
wenzelm
parents:
diff
changeset
|
30 |
( |
33454 | 31 |
type T = spec Item_Net.T; |
33374
8099185908a4
Rules that characterize functional/relational specifications.
wenzelm
parents:
diff
changeset
|
32 |
val empty : T = |
8099185908a4
Rules that characterize functional/relational specifications.
wenzelm
parents:
diff
changeset
|
33 |
Item_Net.init |
33454 | 34 |
(fn ((class1, (ts1, ths1)), (class2, (ts2, ths2))) => |
35 |
class1 = class2 andalso |
|
36 |
eq_list (op aconv) (ts1, ts2) andalso |
|
37 |
eq_list Thm.eq_thm_prop (ths1, ths2)) |
|
38 |
(#1 o #2); |
|
33374
8099185908a4
Rules that characterize functional/relational specifications.
wenzelm
parents:
diff
changeset
|
39 |
val extend = I; |
41472
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
wenzelm
parents:
34948
diff
changeset
|
40 |
val merge = Item_Net.merge; |
33374
8099185908a4
Rules that characterize functional/relational specifications.
wenzelm
parents:
diff
changeset
|
41 |
); |
8099185908a4
Rules that characterize functional/relational specifications.
wenzelm
parents:
diff
changeset
|
42 |
|
61060 | 43 |
|
44 |
(* get *) |
|
45 |
||
46 |
fun get_generic context = |
|
47 |
let |
|
48 |
val thy = Context.theory_of context; |
|
49 |
val transfer = Global_Theory.transfer_theories thy; |
|
50 |
in Item_Net.content (Rules.get context) |> (map o apsnd o apsnd o map) transfer end; |
|
51 |
||
52 |
val get = get_generic o Context.Proof; |
|
53 |
val get_global = get_generic o Context.Theory; |
|
33374
8099185908a4
Rules that characterize functional/relational specifications.
wenzelm
parents:
diff
changeset
|
54 |
|
61060 | 55 |
|
56 |
(* retrieve *) |
|
57 |
||
58 |
fun retrieve_generic context = |
|
59 |
Item_Net.retrieve (Rules.get context) |
|
60 |
#> (map o apsnd o apsnd o map) (Thm.transfer (Context.theory_of context)); |
|
61 |
||
62 |
val retrieve = retrieve_generic o Context.Proof; |
|
63 |
val retrieve_global = retrieve_generic o Context.Theory; |
|
64 |
||
65 |
||
66 |
(* add *) |
|
34948
2d5f2a9f7601
refactoring the predicate compiler; adding theories for Sequences; adding retrieval to Spec_Rules; adding timing to Quickcheck
bulwahn
parents:
33702
diff
changeset
|
67 |
|
33702 | 68 |
fun add class (ts, ths) lthy = |
69 |
let |
|
59621
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents:
59573
diff
changeset
|
70 |
val cts = map (Thm.cterm_of lthy) ts; |
33702 | 71 |
in |
45291
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
wenzelm
parents:
42360
diff
changeset
|
72 |
lthy |> Local_Theory.declaration {syntax = false, pervasive = true} |
33702 | 73 |
(fn phi => |
74 |
let |
|
75 |
val (ts', ths') = |
|
76 |
Morphism.fact phi (map Drule.mk_term cts @ ths) |
|
77 |
|> chop (length cts) |
|
61060 | 78 |
|>> map (Thm.term_of o Drule.dest_term) |
79 |
||> map Thm.trim_context; |
|
33702 | 80 |
in Rules.map (Item_Net.update (class, (ts', ths'))) end) |
81 |
end; |
|
33374
8099185908a4
Rules that characterize functional/relational specifications.
wenzelm
parents:
diff
changeset
|
82 |
|
33454 | 83 |
fun add_global class spec = |
61060 | 84 |
Context.theory_map (Rules.map (Item_Net.update (class, (apsnd o map) Thm.trim_context spec))); |
33374
8099185908a4
Rules that characterize functional/relational specifications.
wenzelm
parents:
diff
changeset
|
85 |
|
8099185908a4
Rules that characterize functional/relational specifications.
wenzelm
parents:
diff
changeset
|
86 |
end; |