| author | paulson <lp15@cam.ac.uk> | 
| Thu, 05 May 2022 16:39:48 +0100 | |
| changeset 75449 | 51e05af57455 | 
| parent 74561 | 8e6c973003c8 | 
| child 78052 | 92d487a28545 | 
| 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  | 
|
| 71202 | 4  | 
Rules that characterize specifications, with optional name and  | 
5  | 
rough classification.  | 
|
6  | 
||
| 33454 | 7  | 
NB: In the face of arbitrary morphisms, the original shape of  | 
8  | 
specifications may get lost.  | 
|
| 
33374
 
8099185908a4
Rules that characterize functional/relational specifications.
 
wenzelm 
parents:  
diff
changeset
 | 
9  | 
*)  | 
| 
 
8099185908a4
Rules that characterize functional/relational specifications.
 
wenzelm 
parents:  
diff
changeset
 | 
10  | 
|
| 
 
8099185908a4
Rules that characterize functional/relational specifications.
 
wenzelm 
parents:  
diff
changeset
 | 
11  | 
signature SPEC_RULES =  | 
| 
 
8099185908a4
Rules that characterize functional/relational specifications.
 
wenzelm 
parents:  
diff
changeset
 | 
12  | 
sig  | 
| 
69996
 
8f2d3a27aff0
more informative Spec_Rules.Equational: support corecursion;
 
wenzelm 
parents: 
69992 
diff
changeset
 | 
13  | 
datatype recursion =  | 
| 
 
8f2d3a27aff0
more informative Spec_Rules.Equational: support corecursion;
 
wenzelm 
parents: 
69992 
diff
changeset
 | 
14  | 
Primrec of string list | Recdef | Primcorec of string list | Corec | Unknown_Recursion  | 
| 70586 | 15  | 
val recursion_ord: recursion ord  | 
| 71202 | 16  | 
val encode_recursion: recursion XML.Encode.T  | 
| 
69992
 
bd3c10813cc4
more informative Spec_Rules.Equational, notably primrec argument types;
 
wenzelm 
parents: 
69991 
diff
changeset
 | 
17  | 
datatype rough_classification = Equational of recursion | Inductive | Co_Inductive | Unknown  | 
| 70586 | 18  | 
val rough_classification_ord: rough_classification ord  | 
| 
69992
 
bd3c10813cc4
more informative Spec_Rules.Equational, notably primrec argument types;
 
wenzelm 
parents: 
69991 
diff
changeset
 | 
19  | 
val equational_primrec: string list -> rough_classification  | 
| 
 
bd3c10813cc4
more informative Spec_Rules.Equational, notably primrec argument types;
 
wenzelm 
parents: 
69991 
diff
changeset
 | 
20  | 
val equational_recdef: rough_classification  | 
| 
69996
 
8f2d3a27aff0
more informative Spec_Rules.Equational: support corecursion;
 
wenzelm 
parents: 
69992 
diff
changeset
 | 
21  | 
val equational_primcorec: string list -> rough_classification  | 
| 
 
8f2d3a27aff0
more informative Spec_Rules.Equational: support corecursion;
 
wenzelm 
parents: 
69992 
diff
changeset
 | 
22  | 
val equational_corec: rough_classification  | 
| 
69992
 
bd3c10813cc4
more informative Spec_Rules.Equational, notably primrec argument types;
 
wenzelm 
parents: 
69991 
diff
changeset
 | 
23  | 
val equational: rough_classification  | 
| 
69991
 
6b097aeb3650
clarified signature: avoid direct comparison on type rough_classification;
 
wenzelm 
parents: 
67649 
diff
changeset
 | 
24  | 
val is_equational: rough_classification -> bool  | 
| 
 
6b097aeb3650
clarified signature: avoid direct comparison on type rough_classification;
 
wenzelm 
parents: 
67649 
diff
changeset
 | 
25  | 
val is_inductive: rough_classification -> bool  | 
| 
 
6b097aeb3650
clarified signature: avoid direct comparison on type rough_classification;
 
wenzelm 
parents: 
67649 
diff
changeset
 | 
26  | 
val is_co_inductive: rough_classification -> bool  | 
| 71179 | 27  | 
val is_relational: rough_classification -> bool  | 
| 
69991
 
6b097aeb3650
clarified signature: avoid direct comparison on type rough_classification;
 
wenzelm 
parents: 
67649 
diff
changeset
 | 
28  | 
val is_unknown: rough_classification -> bool  | 
| 71202 | 29  | 
val encode_rough_classification: rough_classification XML.Encode.T  | 
| 
71210
 
66fa99c85095
tuned signature -- following Export_Theory.Spec_Rule in Scala;
 
wenzelm 
parents: 
71207 
diff
changeset
 | 
30  | 
type spec_rule =  | 
| 
71207
 
8af82f3e03c9
formal position for spec rule (not significant for equality);
 
wenzelm 
parents: 
71202 
diff
changeset
 | 
31  | 
   {pos: Position.T,
 | 
| 
 
8af82f3e03c9
formal position for spec rule (not significant for equality);
 
wenzelm 
parents: 
71202 
diff
changeset
 | 
32  | 
name: string,  | 
| 
 
8af82f3e03c9
formal position for spec rule (not significant for equality);
 
wenzelm 
parents: 
71202 
diff
changeset
 | 
33  | 
rough_classification: rough_classification,  | 
| 
 
8af82f3e03c9
formal position for spec rule (not significant for equality);
 
wenzelm 
parents: 
71202 
diff
changeset
 | 
34  | 
terms: term list,  | 
| 
 
8af82f3e03c9
formal position for spec rule (not significant for equality);
 
wenzelm 
parents: 
71202 
diff
changeset
 | 
35  | 
rules: thm list}  | 
| 
71210
 
66fa99c85095
tuned signature -- following Export_Theory.Spec_Rule in Scala;
 
wenzelm 
parents: 
71207 
diff
changeset
 | 
36  | 
val get: Proof.context -> spec_rule list  | 
| 
 
66fa99c85095
tuned signature -- following Export_Theory.Spec_Rule in Scala;
 
wenzelm 
parents: 
71207 
diff
changeset
 | 
37  | 
val get_global: theory -> spec_rule list  | 
| 
 
66fa99c85095
tuned signature -- following Export_Theory.Spec_Rule in Scala;
 
wenzelm 
parents: 
71207 
diff
changeset
 | 
38  | 
val dest_theory: theory -> spec_rule list  | 
| 
 
66fa99c85095
tuned signature -- following Export_Theory.Spec_Rule in Scala;
 
wenzelm 
parents: 
71207 
diff
changeset
 | 
39  | 
val retrieve: Proof.context -> term -> spec_rule list  | 
| 
 
66fa99c85095
tuned signature -- following Export_Theory.Spec_Rule in Scala;
 
wenzelm 
parents: 
71207 
diff
changeset
 | 
40  | 
val retrieve_global: theory -> term -> spec_rule list  | 
| 
71214
 
5727bcc3c47c
proper spec_rule name via naming/binding/Morphism.binding;
 
wenzelm 
parents: 
71210 
diff
changeset
 | 
41  | 
val add: binding -> rough_classification -> term list -> thm list -> local_theory -> local_theory  | 
| 
 
5727bcc3c47c
proper spec_rule name via naming/binding/Morphism.binding;
 
wenzelm 
parents: 
71210 
diff
changeset
 | 
42  | 
val add_global: binding -> rough_classification -> term list -> thm list -> theory -> theory  | 
| 
33374
 
8099185908a4
Rules that characterize functional/relational specifications.
 
wenzelm 
parents:  
diff
changeset
 | 
43  | 
end;  | 
| 
 
8099185908a4
Rules that characterize functional/relational specifications.
 
wenzelm 
parents:  
diff
changeset
 | 
44  | 
|
| 
 
8099185908a4
Rules that characterize functional/relational specifications.
 
wenzelm 
parents:  
diff
changeset
 | 
45  | 
structure Spec_Rules: SPEC_RULES =  | 
| 
 
8099185908a4
Rules that characterize functional/relational specifications.
 
wenzelm 
parents:  
diff
changeset
 | 
46  | 
struct  | 
| 
 
8099185908a4
Rules that characterize functional/relational specifications.
 
wenzelm 
parents:  
diff
changeset
 | 
47  | 
|
| 
69992
 
bd3c10813cc4
more informative Spec_Rules.Equational, notably primrec argument types;
 
wenzelm 
parents: 
69991 
diff
changeset
 | 
48  | 
(* recursion *)  | 
| 
 
bd3c10813cc4
more informative Spec_Rules.Equational, notably primrec argument types;
 
wenzelm 
parents: 
69991 
diff
changeset
 | 
49  | 
|
| 
69996
 
8f2d3a27aff0
more informative Spec_Rules.Equational: support corecursion;
 
wenzelm 
parents: 
69992 
diff
changeset
 | 
50  | 
datatype recursion =  | 
| 
 
8f2d3a27aff0
more informative Spec_Rules.Equational: support corecursion;
 
wenzelm 
parents: 
69992 
diff
changeset
 | 
51  | 
Primrec of string list | Recdef | Primcorec of string list | Corec | Unknown_Recursion;  | 
| 
 
8f2d3a27aff0
more informative Spec_Rules.Equational: support corecursion;
 
wenzelm 
parents: 
69992 
diff
changeset
 | 
52  | 
|
| 
 
8f2d3a27aff0
more informative Spec_Rules.Equational: support corecursion;
 
wenzelm 
parents: 
69992 
diff
changeset
 | 
53  | 
val recursion_index =  | 
| 
 
8f2d3a27aff0
more informative Spec_Rules.Equational: support corecursion;
 
wenzelm 
parents: 
69992 
diff
changeset
 | 
54  | 
fn Primrec _ => 0 | Recdef => 1 | Primcorec _ => 2 | Corec => 3 | Unknown_Recursion => 4;  | 
| 
69992
 
bd3c10813cc4
more informative Spec_Rules.Equational, notably primrec argument types;
 
wenzelm 
parents: 
69991 
diff
changeset
 | 
55  | 
|
| 
 
bd3c10813cc4
more informative Spec_Rules.Equational, notably primrec argument types;
 
wenzelm 
parents: 
69991 
diff
changeset
 | 
56  | 
fun recursion_ord (Primrec Ts1, Primrec Ts2) = list_ord fast_string_ord (Ts1, Ts2)  | 
| 
69996
 
8f2d3a27aff0
more informative Spec_Rules.Equational: support corecursion;
 
wenzelm 
parents: 
69992 
diff
changeset
 | 
57  | 
| recursion_ord (Primcorec Ts1, Primcorec Ts2) = list_ord fast_string_ord (Ts1, Ts2)  | 
| 
 
8f2d3a27aff0
more informative Spec_Rules.Equational: support corecursion;
 
wenzelm 
parents: 
69992 
diff
changeset
 | 
58  | 
| recursion_ord rs = int_ord (apply2 recursion_index rs);  | 
| 
69992
 
bd3c10813cc4
more informative Spec_Rules.Equational, notably primrec argument types;
 
wenzelm 
parents: 
69991 
diff
changeset
 | 
59  | 
|
| 71202 | 60  | 
val encode_recursion =  | 
61  | 
let open XML.Encode in  | 
|
62  | 
variant  | 
|
63  | 
[fn Primrec a => ([], list string a),  | 
|
64  | 
fn Recdef => ([], []),  | 
|
65  | 
fn Primcorec a => ([], list string a),  | 
|
66  | 
fn Corec => ([], []),  | 
|
67  | 
fn Unknown_Recursion => ([], [])]  | 
|
68  | 
end;  | 
|
69  | 
||
| 
69992
 
bd3c10813cc4
more informative Spec_Rules.Equational, notably primrec argument types;
 
wenzelm 
parents: 
69991 
diff
changeset
 | 
70  | 
|
| 
69991
 
6b097aeb3650
clarified signature: avoid direct comparison on type rough_classification;
 
wenzelm 
parents: 
67649 
diff
changeset
 | 
71  | 
(* rough classification *)  | 
| 
 
6b097aeb3650
clarified signature: avoid direct comparison on type rough_classification;
 
wenzelm 
parents: 
67649 
diff
changeset
 | 
72  | 
|
| 
69992
 
bd3c10813cc4
more informative Spec_Rules.Equational, notably primrec argument types;
 
wenzelm 
parents: 
69991 
diff
changeset
 | 
73  | 
datatype rough_classification = Equational of recursion | Inductive | Co_Inductive | Unknown;  | 
| 
 
bd3c10813cc4
more informative Spec_Rules.Equational, notably primrec argument types;
 
wenzelm 
parents: 
69991 
diff
changeset
 | 
74  | 
|
| 
 
bd3c10813cc4
more informative Spec_Rules.Equational, notably primrec argument types;
 
wenzelm 
parents: 
69991 
diff
changeset
 | 
75  | 
fun rough_classification_ord (Equational r1, Equational r2) = recursion_ord (r1, r2)  | 
| 
 
bd3c10813cc4
more informative Spec_Rules.Equational, notably primrec argument types;
 
wenzelm 
parents: 
69991 
diff
changeset
 | 
76  | 
| rough_classification_ord cs =  | 
| 
 
bd3c10813cc4
more informative Spec_Rules.Equational, notably primrec argument types;
 
wenzelm 
parents: 
69991 
diff
changeset
 | 
77  | 
int_ord (apply2 (fn Equational _ => 0 | Inductive => 1 | Co_Inductive => 2 | Unknown => 3) cs);  | 
| 
69991
 
6b097aeb3650
clarified signature: avoid direct comparison on type rough_classification;
 
wenzelm 
parents: 
67649 
diff
changeset
 | 
78  | 
|
| 
69992
 
bd3c10813cc4
more informative Spec_Rules.Equational, notably primrec argument types;
 
wenzelm 
parents: 
69991 
diff
changeset
 | 
79  | 
val equational_primrec = Equational o Primrec;  | 
| 
 
bd3c10813cc4
more informative Spec_Rules.Equational, notably primrec argument types;
 
wenzelm 
parents: 
69991 
diff
changeset
 | 
80  | 
val equational_recdef = Equational Recdef;  | 
| 
69996
 
8f2d3a27aff0
more informative Spec_Rules.Equational: support corecursion;
 
wenzelm 
parents: 
69992 
diff
changeset
 | 
81  | 
val equational_primcorec = Equational o Primcorec;  | 
| 
 
8f2d3a27aff0
more informative Spec_Rules.Equational: support corecursion;
 
wenzelm 
parents: 
69992 
diff
changeset
 | 
82  | 
val equational_corec = Equational Corec;  | 
| 
69992
 
bd3c10813cc4
more informative Spec_Rules.Equational, notably primrec argument types;
 
wenzelm 
parents: 
69991 
diff
changeset
 | 
83  | 
val equational = Equational Unknown_Recursion;  | 
| 
69991
 
6b097aeb3650
clarified signature: avoid direct comparison on type rough_classification;
 
wenzelm 
parents: 
67649 
diff
changeset
 | 
84  | 
|
| 
69992
 
bd3c10813cc4
more informative Spec_Rules.Equational, notably primrec argument types;
 
wenzelm 
parents: 
69991 
diff
changeset
 | 
85  | 
val is_equational = fn Equational _ => true | _ => false;  | 
| 
69991
 
6b097aeb3650
clarified signature: avoid direct comparison on type rough_classification;
 
wenzelm 
parents: 
67649 
diff
changeset
 | 
86  | 
val is_inductive = fn Inductive => true | _ => false;  | 
| 
 
6b097aeb3650
clarified signature: avoid direct comparison on type rough_classification;
 
wenzelm 
parents: 
67649 
diff
changeset
 | 
87  | 
val is_co_inductive = fn Co_Inductive => true | _ => false;  | 
| 71179 | 88  | 
val is_relational = is_inductive orf is_co_inductive;  | 
| 
69991
 
6b097aeb3650
clarified signature: avoid direct comparison on type rough_classification;
 
wenzelm 
parents: 
67649 
diff
changeset
 | 
89  | 
val is_unknown = fn Unknown => true | _ => false;  | 
| 
 
6b097aeb3650
clarified signature: avoid direct comparison on type rough_classification;
 
wenzelm 
parents: 
67649 
diff
changeset
 | 
90  | 
|
| 71202 | 91  | 
val encode_rough_classification =  | 
92  | 
let open XML.Encode in  | 
|
93  | 
variant  | 
|
94  | 
[fn Equational r => ([], encode_recursion r),  | 
|
95  | 
fn Inductive => ([], []),  | 
|
96  | 
fn Co_Inductive => ([], []),  | 
|
97  | 
fn Unknown => ([], [])]  | 
|
98  | 
end;  | 
|
99  | 
||
| 
69991
 
6b097aeb3650
clarified signature: avoid direct comparison on type rough_classification;
 
wenzelm 
parents: 
67649 
diff
changeset
 | 
100  | 
|
| 61060 | 101  | 
(* rules *)  | 
| 
33374
 
8099185908a4
Rules that characterize functional/relational specifications.
 
wenzelm 
parents:  
diff
changeset
 | 
102  | 
|
| 
71210
 
66fa99c85095
tuned signature -- following Export_Theory.Spec_Rule in Scala;
 
wenzelm 
parents: 
71207 
diff
changeset
 | 
103  | 
type spec_rule =  | 
| 
71207
 
8af82f3e03c9
formal position for spec rule (not significant for equality);
 
wenzelm 
parents: 
71202 
diff
changeset
 | 
104  | 
 {pos: Position.T,
 | 
| 
 
8af82f3e03c9
formal position for spec rule (not significant for equality);
 
wenzelm 
parents: 
71202 
diff
changeset
 | 
105  | 
name: string,  | 
| 
 
8af82f3e03c9
formal position for spec rule (not significant for equality);
 
wenzelm 
parents: 
71202 
diff
changeset
 | 
106  | 
rough_classification: rough_classification,  | 
| 
 
8af82f3e03c9
formal position for spec rule (not significant for equality);
 
wenzelm 
parents: 
71202 
diff
changeset
 | 
107  | 
terms: term list,  | 
| 
 
8af82f3e03c9
formal position for spec rule (not significant for equality);
 
wenzelm 
parents: 
71202 
diff
changeset
 | 
108  | 
rules: thm list};  | 
| 71179 | 109  | 
|
| 
71210
 
66fa99c85095
tuned signature -- following Export_Theory.Spec_Rule in Scala;
 
wenzelm 
parents: 
71207 
diff
changeset
 | 
110  | 
fun eq_spec (specs: spec_rule * spec_rule) =  | 
| 71179 | 111  | 
(op =) (apply2 #name specs) andalso  | 
112  | 
is_equal (rough_classification_ord (apply2 #rough_classification specs)) andalso  | 
|
113  | 
eq_list (op aconv) (apply2 #terms specs) andalso  | 
|
114  | 
eq_list Thm.eq_thm_prop (apply2 #rules specs);  | 
|
115  | 
||
| 
71210
 
66fa99c85095
tuned signature -- following Export_Theory.Spec_Rule in Scala;
 
wenzelm 
parents: 
71207 
diff
changeset
 | 
116  | 
fun map_spec_rules f ({pos, name, rough_classification, terms, rules}: spec_rule) : spec_rule =
 | 
| 
71207
 
8af82f3e03c9
formal position for spec rule (not significant for equality);
 
wenzelm 
parents: 
71202 
diff
changeset
 | 
117  | 
  {pos = pos, name = name, rough_classification = rough_classification, terms = terms,
 | 
| 
 
8af82f3e03c9
formal position for spec rule (not significant for equality);
 
wenzelm 
parents: 
71202 
diff
changeset
 | 
118  | 
rules = map f rules};  | 
| 
33374
 
8099185908a4
Rules that characterize functional/relational specifications.
 
wenzelm 
parents:  
diff
changeset
 | 
119  | 
|
| 33519 | 120  | 
structure Rules = Generic_Data  | 
| 
33374
 
8099185908a4
Rules that characterize functional/relational specifications.
 
wenzelm 
parents:  
diff
changeset
 | 
121  | 
(  | 
| 
71210
 
66fa99c85095
tuned signature -- following Export_Theory.Spec_Rule in Scala;
 
wenzelm 
parents: 
71207 
diff
changeset
 | 
122  | 
type T = spec_rule Item_Net.T;  | 
| 71179 | 123  | 
val empty : T = Item_Net.init eq_spec #terms;  | 
| 
41472
 
f6ab14e61604
misc tuning and comments based on review of Theory_Data, Proof_Data, Generic_Data usage;
 
wenzelm 
parents: 
34948 
diff
changeset
 | 
124  | 
val merge = Item_Net.merge;  | 
| 
33374
 
8099185908a4
Rules that characterize functional/relational specifications.
 
wenzelm 
parents:  
diff
changeset
 | 
125  | 
);  | 
| 
 
8099185908a4
Rules that characterize functional/relational specifications.
 
wenzelm 
parents:  
diff
changeset
 | 
126  | 
|
| 61060 | 127  | 
|
128  | 
(* get *)  | 
|
129  | 
||
| 71202 | 130  | 
fun get_generic imports context =  | 
| 61060 | 131  | 
let  | 
132  | 
val thy = Context.theory_of context;  | 
|
133  | 
val transfer = Global_Theory.transfer_theories thy;  | 
|
| 71202 | 134  | 
fun imported spec =  | 
135  | 
imports |> exists (fn thy => Item_Net.member (Rules.get (Context.Theory thy)) spec);  | 
|
| 71179 | 136  | 
in  | 
137  | 
Item_Net.content (Rules.get context)  | 
|
| 71202 | 138  | 
|> filter_out imported  | 
| 71179 | 139  | 
|> (map o map_spec_rules) transfer  | 
140  | 
end;  | 
|
| 61060 | 141  | 
|
| 71202 | 142  | 
val get = get_generic [] o Context.Proof;  | 
143  | 
val get_global = get_generic [] o Context.Theory;  | 
|
144  | 
||
145  | 
fun dest_theory thy = rev (get_generic (Theory.parents_of thy) (Context.Theory thy));  | 
|
| 
33374
 
8099185908a4
Rules that characterize functional/relational specifications.
 
wenzelm 
parents:  
diff
changeset
 | 
146  | 
|
| 61060 | 147  | 
|
148  | 
(* retrieve *)  | 
|
149  | 
||
150  | 
fun retrieve_generic context =  | 
|
151  | 
Item_Net.retrieve (Rules.get context)  | 
|
| 71179 | 152  | 
#> (map o map_spec_rules) (Thm.transfer'' context);  | 
| 61060 | 153  | 
|
154  | 
val retrieve = retrieve_generic o Context.Proof;  | 
|
155  | 
val retrieve_global = retrieve_generic o Context.Theory;  | 
|
156  | 
||
157  | 
||
158  | 
(* 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
 | 
159  | 
|
| 
71214
 
5727bcc3c47c
proper spec_rule name via naming/binding/Morphism.binding;
 
wenzelm 
parents: 
71210 
diff
changeset
 | 
160  | 
fun add b rough_classification terms rules lthy =  | 
| 
71216
 
e64c249d3d98
proper dynamic position of application context, e.g. relevant for 'global_interpretation';
 
wenzelm 
parents: 
71214 
diff
changeset
 | 
161  | 
let val thms0 = map Thm.trim_context (map (Drule.mk_term o Thm.cterm_of lthy) terms @ rules) in  | 
| 
45291
 
57cd50f98fdc
uniform Local_Theory.declaration with explicit params;
 
wenzelm 
parents: 
42360 
diff
changeset
 | 
162  | 
    lthy |> Local_Theory.declaration {syntax = false, pervasive = true}
 | 
| 71179 | 163  | 
(fn phi => fn context =>  | 
| 33702 | 164  | 
let  | 
| 
71216
 
e64c249d3d98
proper dynamic position of application context, e.g. relevant for 'global_interpretation';
 
wenzelm 
parents: 
71214 
diff
changeset
 | 
165  | 
val pos = Position.thread_data ();  | 
| 
71214
 
5727bcc3c47c
proper spec_rule name via naming/binding/Morphism.binding;
 
wenzelm 
parents: 
71210 
diff
changeset
 | 
166  | 
val name = Name_Space.full_name (Name_Space.naming_of context) (Morphism.binding phi b);  | 
| 71179 | 167  | 
val (terms', rules') =  | 
168  | 
map (Thm.transfer (Context.theory_of context)) thms0  | 
|
169  | 
|> Morphism.fact phi  | 
|
170  | 
|> chop (length terms)  | 
|
| 61060 | 171  | 
|>> map (Thm.term_of o Drule.dest_term)  | 
172  | 
||> map Thm.trim_context;  | 
|
| 71179 | 173  | 
in  | 
174  | 
context |> (Rules.map o Item_Net.update)  | 
|
| 
71207
 
8af82f3e03c9
formal position for spec rule (not significant for equality);
 
wenzelm 
parents: 
71202 
diff
changeset
 | 
175  | 
            {pos = pos, name = name, rough_classification = rough_classification,
 | 
| 71179 | 176  | 
terms = terms', rules = rules'}  | 
177  | 
end)  | 
|
| 33702 | 178  | 
end;  | 
| 
33374
 
8099185908a4
Rules that characterize functional/relational specifications.
 
wenzelm 
parents:  
diff
changeset
 | 
179  | 
|
| 
71214
 
5727bcc3c47c
proper spec_rule name via naming/binding/Morphism.binding;
 
wenzelm 
parents: 
71210 
diff
changeset
 | 
180  | 
fun add_global b rough_classification terms rules thy =  | 
| 
 
5727bcc3c47c
proper spec_rule name via naming/binding/Morphism.binding;
 
wenzelm 
parents: 
71210 
diff
changeset
 | 
181  | 
thy |> (Context.theory_map o Rules.map o Item_Net.update)  | 
| 
 
5727bcc3c47c
proper spec_rule name via naming/binding/Morphism.binding;
 
wenzelm 
parents: 
71210 
diff
changeset
 | 
182  | 
   {pos = Position.thread_data (),
 | 
| 
 
5727bcc3c47c
proper spec_rule name via naming/binding/Morphism.binding;
 
wenzelm 
parents: 
71210 
diff
changeset
 | 
183  | 
name = Sign.full_name thy b,  | 
| 
 
5727bcc3c47c
proper spec_rule name via naming/binding/Morphism.binding;
 
wenzelm 
parents: 
71210 
diff
changeset
 | 
184  | 
rough_classification = rough_classification,  | 
| 
 
5727bcc3c47c
proper spec_rule name via naming/binding/Morphism.binding;
 
wenzelm 
parents: 
71210 
diff
changeset
 | 
185  | 
terms = terms,  | 
| 
 
5727bcc3c47c
proper spec_rule name via naming/binding/Morphism.binding;
 
wenzelm 
parents: 
71210 
diff
changeset
 | 
186  | 
rules = map Thm.trim_context rules};  | 
| 
33374
 
8099185908a4
Rules that characterize functional/relational specifications.
 
wenzelm 
parents:  
diff
changeset
 | 
187  | 
|
| 
 
8099185908a4
Rules that characterize functional/relational specifications.
 
wenzelm 
parents:  
diff
changeset
 | 
188  | 
end;  |