| author | Christian Urban <urbanc@in.tum.de> | 
| Sun, 29 Aug 2010 01:45:07 +0800 | |
| changeset 2451 | d2e929f51fa9 | 
| parent 2450 | 217ef3e4282e | 
| permissions | -rw-r--r-- | 
| 1941 | 1  | 
theory NewParser  | 
| 
2435
 
3772bb3bd7ce
cleaning of unused files and code
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
2  | 
imports  | 
| 
 
3772bb3bd7ce
cleaning of unused files and code
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
3  | 
"../Nominal-General/Nominal2_Base"  | 
| 
 
3772bb3bd7ce
cleaning of unused files and code
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
4  | 
"../Nominal-General/Nominal2_Eqvt"  | 
| 
 
3772bb3bd7ce
cleaning of unused files and code
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
5  | 
"../Nominal-General/Nominal2_Supp"  | 
| 
 
3772bb3bd7ce
cleaning of unused files and code
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
6  | 
"Nominal2_FSet"  | 
| 
 
3772bb3bd7ce
cleaning of unused files and code
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
7  | 
"Abs"  | 
| 
 
3772bb3bd7ce
cleaning of unused files and code
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
8  | 
uses ("nominal_dt_rawperm.ML")
 | 
| 
 
3772bb3bd7ce
cleaning of unused files and code
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
9  | 
     ("nominal_dt_rawfuns.ML")
 | 
| 
 
3772bb3bd7ce
cleaning of unused files and code
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
10  | 
     ("nominal_dt_alpha.ML")
 | 
| 
 
3772bb3bd7ce
cleaning of unused files and code
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
11  | 
     ("nominal_dt_quot.ML")
 | 
| 2448 | 12  | 
     ("nominal_dt_supp.ML")
 | 
| 1941 | 13  | 
begin  | 
14  | 
||
| 
2435
 
3772bb3bd7ce
cleaning of unused files and code
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
15  | 
use "nominal_dt_rawperm.ML"  | 
| 
 
3772bb3bd7ce
cleaning of unused files and code
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
16  | 
ML {* open Nominal_Dt_RawPerm *}
 | 
| 
 
3772bb3bd7ce
cleaning of unused files and code
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
17  | 
|
| 
 
3772bb3bd7ce
cleaning of unused files and code
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
18  | 
use "nominal_dt_rawfuns.ML"  | 
| 
 
3772bb3bd7ce
cleaning of unused files and code
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
19  | 
ML {* open Nominal_Dt_RawFuns *}
 | 
| 
 
3772bb3bd7ce
cleaning of unused files and code
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
20  | 
|
| 
 
3772bb3bd7ce
cleaning of unused files and code
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
21  | 
use "nominal_dt_alpha.ML"  | 
| 
 
3772bb3bd7ce
cleaning of unused files and code
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
22  | 
ML {* open Nominal_Dt_Alpha *}
 | 
| 
 
3772bb3bd7ce
cleaning of unused files and code
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
23  | 
|
| 
 
3772bb3bd7ce
cleaning of unused files and code
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
24  | 
use "nominal_dt_quot.ML"  | 
| 
 
3772bb3bd7ce
cleaning of unused files and code
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
25  | 
ML {* open Nominal_Dt_Quot *}
 | 
| 
 
3772bb3bd7ce
cleaning of unused files and code
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
26  | 
|
| 2448 | 27  | 
use "nominal_dt_supp.ML"  | 
28  | 
ML {* open Nominal_Dt_Supp *}
 | 
|
| 
2288
 
3b83960f9544
new fv/fv_bn function (supp breaks now); exported raw perms and raw funs into separate ML-files
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2146 
diff
changeset
 | 
29  | 
|
| 1941 | 30  | 
section{* Interface for nominal_datatype *}
 | 
31  | 
||
| 2398 | 32  | 
ML {*
 | 
33  | 
(* attributes *)  | 
|
| 2448 | 34  | 
val eqvt_attr = Attrib.internal (K Nominal_ThmDecls.eqvt_add)  | 
35  | 
val rsp_attr = Attrib.internal (K Quotient_Info.rsp_rules_add)  | 
|
36  | 
val simp_attr = Attrib.internal (K Simplifier.simp_add)  | 
|
| 2398 | 37  | 
|
38  | 
*}  | 
|
39  | 
||
| 
2424
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2411 
diff
changeset
 | 
40  | 
ML {* print_depth 50 *}
 | 
| 1941 | 41  | 
|
42  | 
ML {*
 | 
|
43  | 
fun get_cnstrs dts =  | 
|
44  | 
map (fn (_, _, _, constrs) => constrs) dts  | 
|
45  | 
||
46  | 
fun get_typed_cnstrs dts =  | 
|
47  | 
flat (map (fn (_, bn, _, constrs) =>  | 
|
48  | 
(map (fn (bn', _, _) => (Binding.name_of bn, Binding.name_of bn')) constrs)) dts)  | 
|
49  | 
||
50  | 
fun get_cnstr_strs dts =  | 
|
51  | 
map (fn (bn, _, _) => Binding.name_of bn) (flat (get_cnstrs dts))  | 
|
52  | 
||
53  | 
fun get_bn_fun_strs bn_funs =  | 
|
54  | 
map (fn (bn_fun, _, _) => Binding.name_of bn_fun) bn_funs  | 
|
55  | 
*}  | 
|
56  | 
||
| 
2106
 
409ecb7284dd
properly exported defined bn-functions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2105 
diff
changeset
 | 
57  | 
|
| 1944 | 58  | 
text {* Infrastructure for adding "_raw" to types and terms *}
 | 
59  | 
||
| 1941 | 60  | 
ML {*
 | 
61  | 
fun add_raw s = s ^ "_raw"  | 
|
62  | 
fun add_raws ss = map add_raw ss  | 
|
63  | 
fun raw_bind bn = Binding.suffix_name "_raw" bn  | 
|
64  | 
||
65  | 
fun replace_str ss s =  | 
|
66  | 
case (AList.lookup (op=) ss s) of  | 
|
67  | 
SOME s' => s'  | 
|
68  | 
| NONE => s  | 
|
69  | 
||
70  | 
fun replace_typ ty_ss (Type (a, Ts)) = Type (replace_str ty_ss a, map (replace_typ ty_ss) Ts)  | 
|
71  | 
| replace_typ ty_ss T = T  | 
|
72  | 
||
73  | 
fun raw_dts ty_ss dts =  | 
|
74  | 
let  | 
|
75  | 
fun raw_dts_aux1 (bind, tys, mx) =  | 
|
76  | 
(raw_bind bind, map (replace_typ ty_ss) tys, mx)  | 
|
77  | 
||
78  | 
fun raw_dts_aux2 (ty_args, bind, mx, constrs) =  | 
|
79  | 
(ty_args, raw_bind bind, mx, map raw_dts_aux1 constrs)  | 
|
80  | 
in  | 
|
81  | 
map raw_dts_aux2 dts  | 
|
82  | 
end  | 
|
83  | 
||
84  | 
fun replace_aterm trm_ss (Const (a, T)) = Const (replace_str trm_ss a, T)  | 
|
85  | 
| replace_aterm trm_ss (Free (a, T)) = Free (replace_str trm_ss a, T)  | 
|
86  | 
| replace_aterm trm_ss trm = trm  | 
|
87  | 
||
88  | 
fun replace_term trm_ss ty_ss trm =  | 
|
89  | 
trm |> Term.map_aterms (replace_aterm trm_ss) |> map_types (replace_typ ty_ss)  | 
|
90  | 
*}  | 
|
91  | 
||
92  | 
ML {*
 | 
|
93  | 
fun rawify_dts dt_names dts dts_env =  | 
|
94  | 
let  | 
|
95  | 
val raw_dts = raw_dts dts_env dts  | 
|
96  | 
val raw_dt_names = add_raws dt_names  | 
|
97  | 
in  | 
|
98  | 
(raw_dt_names, raw_dts)  | 
|
99  | 
end  | 
|
100  | 
*}  | 
|
101  | 
||
102  | 
ML {*
 | 
|
103  | 
fun rawify_bn_funs dts_env cnstrs_env bn_fun_env bn_funs bn_eqs =  | 
|
104  | 
let  | 
|
105  | 
val bn_funs' = map (fn (bn, ty, mx) =>  | 
|
| 
2304
 
8a98171ba1fc
all raw definitions are defined using function
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2302 
diff
changeset
 | 
106  | 
(raw_bind bn, SOME (replace_typ dts_env ty), mx)) bn_funs  | 
| 1941 | 107  | 
|
108  | 
val bn_eqs' = map (fn (attr, trm) =>  | 
|
109  | 
(attr, replace_term (cnstrs_env @ bn_fun_env) dts_env trm)) bn_eqs  | 
|
110  | 
in  | 
|
111  | 
(bn_funs', bn_eqs')  | 
|
112  | 
end  | 
|
113  | 
*}  | 
|
114  | 
||
115  | 
ML {* 
 | 
|
116  | 
fun rawify_bclauses dts_env cnstrs_env bn_fun_env bclauses =  | 
|
117  | 
let  | 
|
118  | 
fun rawify_bnds bnds =  | 
|
119  | 
map (apfst (Option.map (replace_term (cnstrs_env @ bn_fun_env) dts_env))) bnds  | 
|
120  | 
||
| 
2288
 
3b83960f9544
new fv/fv_bn function (supp breaks now); exported raw perms and raw funs into separate ML-files
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2146 
diff
changeset
 | 
121  | 
fun rawify_bclause (BC (mode, bnds, bdys)) = BC (mode, rawify_bnds bnds, bdys)  | 
| 1941 | 122  | 
in  | 
123  | 
map (map (map rawify_bclause)) bclauses  | 
|
124  | 
end  | 
|
125  | 
*}  | 
|
126  | 
||
| 
2143
 
871d8a5e0c67
somewhat simplified the main parsing function; failed to move a Note-statement to define_raw_perms
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2142 
diff
changeset
 | 
127  | 
(* strip_bn_fun takes a rhs of a bn function: this can only contain unions or  | 
| 
 
871d8a5e0c67
somewhat simplified the main parsing function; failed to move a Note-statement to define_raw_perms
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2142 
diff
changeset
 | 
128  | 
appends of elements; in case of recursive calls it retruns also the applied  | 
| 
 
871d8a5e0c67
somewhat simplified the main parsing function; failed to move a Note-statement to define_raw_perms
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2142 
diff
changeset
 | 
129  | 
bn function *)  | 
| 1941 | 130  | 
ML {*
 | 
| 
2294
 
72ad4e766acf
properly exported bn_descr
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2293 
diff
changeset
 | 
131  | 
fun strip_bn_fun lthy args t =  | 
| 
 
72ad4e766acf
properly exported bn_descr
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2293 
diff
changeset
 | 
132  | 
let  | 
| 
 
72ad4e766acf
properly exported bn_descr
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2293 
diff
changeset
 | 
133  | 
fun aux t =  | 
| 
 
72ad4e766acf
properly exported bn_descr
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2293 
diff
changeset
 | 
134  | 
case t of  | 
| 
 
72ad4e766acf
properly exported bn_descr
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2293 
diff
changeset
 | 
135  | 
      Const (@{const_name sup}, _) $ l $ r => aux l @ aux r
 | 
| 
 
72ad4e766acf
properly exported bn_descr
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2293 
diff
changeset
 | 
136  | 
    | Const (@{const_name append}, _) $ l $ r => aux l @ aux r
 | 
| 
 
72ad4e766acf
properly exported bn_descr
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2293 
diff
changeset
 | 
137  | 
    | Const (@{const_name insert}, _) $ (Const (@{const_name atom}, _) $ (x as Var _)) $ y =>
 | 
| 
 
72ad4e766acf
properly exported bn_descr
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2293 
diff
changeset
 | 
138  | 
(find_index (equal x) args, NONE) :: aux y  | 
| 
 
72ad4e766acf
properly exported bn_descr
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2293 
diff
changeset
 | 
139  | 
    | Const (@{const_name Cons}, _) $ (Const (@{const_name atom}, _) $ (x as Var _)) $ y =>
 | 
| 
 
72ad4e766acf
properly exported bn_descr
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2293 
diff
changeset
 | 
140  | 
(find_index (equal x) args, NONE) :: aux y  | 
| 
 
72ad4e766acf
properly exported bn_descr
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2293 
diff
changeset
 | 
141  | 
    | Const (@{const_name bot}, _) => []
 | 
| 
 
72ad4e766acf
properly exported bn_descr
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2293 
diff
changeset
 | 
142  | 
    | Const (@{const_name Nil}, _) => []
 | 
| 
 
72ad4e766acf
properly exported bn_descr
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2293 
diff
changeset
 | 
143  | 
| (f as Const _) $ (x as Var _) => [(find_index (equal x) args, SOME f)]  | 
| 
 
72ad4e766acf
properly exported bn_descr
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2293 
diff
changeset
 | 
144  | 
    | _ => error ("Unsupported binding function: " ^ (Syntax.string_of_term lthy t))
 | 
| 
 
72ad4e766acf
properly exported bn_descr
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2293 
diff
changeset
 | 
145  | 
in  | 
| 
 
72ad4e766acf
properly exported bn_descr
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2293 
diff
changeset
 | 
146  | 
aux t  | 
| 
 
72ad4e766acf
properly exported bn_descr
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2293 
diff
changeset
 | 
147  | 
end  | 
| 1941 | 148  | 
*}  | 
149  | 
||
150  | 
ML {*
 | 
|
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
151  | 
(** definition of the raw binding functions **)  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
152  | 
|
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
153  | 
(* TODO: needs cleaning *)  | 
| 1941 | 154  | 
fun find [] _ = error ("cannot find element")
 | 
155  | 
| find ((x, z)::xs) y = if (Long_Name.base_name x) = y then z else find xs y  | 
|
156  | 
||
| 
2295
 
8aff3f3ce47f
started to work on alpha
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2294 
diff
changeset
 | 
157  | 
fun prep_bn_info lthy dt_names dts eqs =  | 
| 1941 | 158  | 
let  | 
159  | 
fun aux eq =  | 
|
160  | 
let  | 
|
161  | 
val (lhs, rhs) = eq  | 
|
162  | 
|> HOLogic.dest_Trueprop  | 
|
163  | 
|> HOLogic.dest_eq  | 
|
164  | 
val (bn_fun, [cnstr]) = strip_comb lhs  | 
|
| 
2294
 
72ad4e766acf
properly exported bn_descr
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2293 
diff
changeset
 | 
165  | 
val (_, ty) = dest_Const bn_fun  | 
| 1941 | 166  | 
val (ty_name, _) = dest_Type (domain_type ty)  | 
167  | 
val dt_index = find_index (fn x => x = ty_name) dt_names  | 
|
| 
2294
 
72ad4e766acf
properly exported bn_descr
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2293 
diff
changeset
 | 
168  | 
val (cnstr_head, cnstr_args) = strip_comb cnstr  | 
| 
 
72ad4e766acf
properly exported bn_descr
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2293 
diff
changeset
 | 
169  | 
val rhs_elements = strip_bn_fun lthy cnstr_args rhs  | 
| 1941 | 170  | 
in  | 
| 
2308
 
387fcbd33820
fixed problem with bn_info
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2306 
diff
changeset
 | 
171  | 
(dt_index, (bn_fun, (cnstr_head, rhs_elements)))  | 
| 
2122
 
24ca435ead14
added term4 back to the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2119 
diff
changeset
 | 
172  | 
end  | 
| 1941 | 173  | 
fun order dts i ts =  | 
174  | 
let  | 
|
175  | 
val dt = nth dts i  | 
|
176  | 
val cts = map (fn (x, _, _) => Binding.name_of x) ((fn (_, _, _, x) => x) dt)  | 
|
177  | 
val ts' = map (fn (x, y) => (fst (dest_Const x), y)) ts  | 
|
178  | 
in  | 
|
179  | 
map (find ts') cts  | 
|
180  | 
end  | 
|
181  | 
||
182  | 
val unordered = AList.group (op=) (map aux eqs)  | 
|
183  | 
val unordered' = map (fn (x, y) => (x, AList.group (op=) y)) unordered  | 
|
184  | 
val ordered = map (fn (x, y) => (x, map (fn (v, z) => (v, order dts x z)) y)) unordered'  | 
|
| 
2118
 
0e52851acac4
moved the data-transformation into the parser
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2107 
diff
changeset
 | 
185  | 
val ordered' = flat (map (fn (ith, l) => map (fn (bn, data) => (bn, ith, data)) l) ordered)  | 
| 
2122
 
24ca435ead14
added term4 back to the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2119 
diff
changeset
 | 
186  | 
|
| 
2288
 
3b83960f9544
new fv/fv_bn function (supp breaks now); exported raw perms and raw funs into separate ML-files
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2146 
diff
changeset
 | 
187  | 
  (*val _ = tracing ("eqs\n" ^ cat_lines (map (Syntax.string_of_term lthy) eqs))*)
 | 
| 
2142
 
c39d4fe31100
moved the exporting part into the parser (this is still a hack); re-added CoreHaskell again to the examples - there seems to be a problem with the variable name pat
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2125 
diff
changeset
 | 
188  | 
  (*val _ = tracing ("map eqs\n" ^ @{make_string} (map aux2 eqs))*)
 | 
| 
2288
 
3b83960f9544
new fv/fv_bn function (supp breaks now); exported raw perms and raw funs into separate ML-files
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2146 
diff
changeset
 | 
189  | 
  (*val _ = tracing ("ordered'\n" ^ @{make_string} ordered')*)
 | 
| 1941 | 190  | 
in  | 
| 
2125
 
60ee289a8c63
made out of STEPS a configuration value so that it can be set individually in each file
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2122 
diff
changeset
 | 
191  | 
ordered'  | 
| 1941 | 192  | 
end  | 
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
193  | 
|
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
194  | 
|
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
195  | 
fun define_raw_bns dt_names dts raw_bn_funs raw_bn_eqs constr_thms size_thms lthy =  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
196  | 
if null raw_bn_funs  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
197  | 
then ([], [], [], [], lthy)  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
198  | 
else  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
199  | 
let  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
200  | 
val (_, lthy1) = Function.add_function raw_bn_funs raw_bn_eqs  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
201  | 
Function_Common.default_config (pat_completeness_simp constr_thms) lthy  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
202  | 
|
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
203  | 
val (info, lthy2) = prove_termination size_thms (Local_Theory.restore lthy1)  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
204  | 
      val {fs, simps, inducts, ...} = info
 | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
205  | 
|
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
206  | 
val raw_bn_induct = (the inducts)  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
207  | 
val raw_bn_eqs = the simps  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
208  | 
|
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
209  | 
val raw_bn_info =  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
210  | 
prep_bn_info lthy dt_names dts (map prop_of raw_bn_eqs)  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
211  | 
in  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
212  | 
(fs, raw_bn_eqs, raw_bn_info, raw_bn_induct, lthy2)  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
213  | 
end  | 
| 1941 | 214  | 
*}  | 
215  | 
||
216  | 
ML {*
 | 
|
| 
2410
 
2bbdb9c427b5
improved runtime slightly, by constructing an explicit size measure for the function definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2409 
diff
changeset
 | 
217  | 
fun define_raw_dts dts bn_funs bn_eqs binds lthy =  | 
| 1941 | 218  | 
let  | 
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
219  | 
val thy = Local_Theory.exit_global lthy  | 
| 1941 | 220  | 
val thy_name = Context.theory_name thy  | 
221  | 
||
222  | 
val dt_names = map (fn (_, s, _, _) => Binding.name_of s) dts  | 
|
223  | 
val dt_full_names = map (Long_Name.qualify thy_name) dt_names  | 
|
224  | 
val dt_full_names' = add_raws dt_full_names  | 
|
225  | 
val dts_env = dt_full_names ~~ dt_full_names'  | 
|
226  | 
||
227  | 
val cnstrs = get_cnstr_strs dts  | 
|
228  | 
val cnstrs_ty = get_typed_cnstrs dts  | 
|
229  | 
val cnstrs_full_names = map (Long_Name.qualify thy_name) cnstrs  | 
|
230  | 
val cnstrs_full_names' = map (fn (x, y) => Long_Name.qualify thy_name  | 
|
231  | 
(Long_Name.qualify (add_raw x) (add_raw y))) cnstrs_ty  | 
|
232  | 
val cnstrs_env = cnstrs_full_names ~~ cnstrs_full_names'  | 
|
233  | 
||
234  | 
val bn_fun_strs = get_bn_fun_strs bn_funs  | 
|
235  | 
val bn_fun_strs' = add_raws bn_fun_strs  | 
|
236  | 
val bn_fun_env = bn_fun_strs ~~ bn_fun_strs'  | 
|
237  | 
val bn_fun_full_env = map (pairself (Long_Name.qualify thy_name))  | 
|
238  | 
(bn_fun_strs ~~ bn_fun_strs')  | 
|
239  | 
||
240  | 
val (raw_dt_names, raw_dts) = rawify_dts dt_names dts dts_env  | 
|
241  | 
val (raw_bn_funs, raw_bn_eqs) = rawify_bn_funs dts_env cnstrs_env bn_fun_env bn_funs bn_eqs  | 
|
242  | 
val raw_bclauses = rawify_bclauses dts_env cnstrs_env bn_fun_full_env binds  | 
|
243  | 
||
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
244  | 
val (raw_dt_full_names, thy1) =  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
245  | 
Datatype.add_datatype Datatype.default_config raw_dt_names raw_dts thy  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
246  | 
|
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
247  | 
val lthy1 = Named_Target.theory_init thy1  | 
| 
2304
 
8a98171ba1fc
all raw definitions are defined using function
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2302 
diff
changeset
 | 
248  | 
in  | 
| 
2337
 
b151399bd2c3
fixed according to changes in quotient
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2336 
diff
changeset
 | 
249  | 
(raw_dt_full_names, raw_dts, raw_bclauses, raw_bn_funs, raw_bn_eqs, lthy1)  | 
| 
2304
 
8a98171ba1fc
all raw definitions are defined using function
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2302 
diff
changeset
 | 
250  | 
end  | 
| 
 
8a98171ba1fc
all raw definitions are defined using function
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2302 
diff
changeset
 | 
251  | 
*}  | 
| 
 
8a98171ba1fc
all raw definitions are defined using function
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2302 
diff
changeset
 | 
252  | 
|
| 
 
8a98171ba1fc
all raw definitions are defined using function
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2302 
diff
changeset
 | 
253  | 
|
| 
2046
 
73c50e913db6
tuned and added some comments to the code; added also an exception for early exit of the nominal2_cmd function
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2037 
diff
changeset
 | 
254  | 
ML {*
 | 
| 
 
73c50e913db6
tuned and added some comments to the code; added also an exception for early exit of the nominal2_cmd function
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2037 
diff
changeset
 | 
255  | 
(* for testing porposes - to exit the procedure early *)  | 
| 
 
73c50e913db6
tuned and added some comments to the code; added also an exception for early exit of the nominal2_cmd function
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2037 
diff
changeset
 | 
256  | 
exception TEST of Proof.context  | 
| 
 
73c50e913db6
tuned and added some comments to the code; added also an exception for early exit of the nominal2_cmd function
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2037 
diff
changeset
 | 
257  | 
|
| 
2436
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2435 
diff
changeset
 | 
258  | 
val (STEPS, STEPS_setup) = Attrib.config_int "STEPS" (K 100);  | 
| 
2125
 
60ee289a8c63
made out of STEPS a configuration value so that it can be set individually in each file
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2122 
diff
changeset
 | 
259  | 
|
| 
 
60ee289a8c63
made out of STEPS a configuration value so that it can be set individually in each file
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2122 
diff
changeset
 | 
260  | 
fun get_STEPS ctxt = Config.get ctxt STEPS  | 
| 
2046
 
73c50e913db6
tuned and added some comments to the code; added also an exception for early exit of the nominal2_cmd function
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2037 
diff
changeset
 | 
261  | 
*}  | 
| 
2008
 
1bddffddc03f
attempted to remove dependency on (old) Fv and (old) Parser; lifting still uses Fv.thy; the examples do not work at the moment (with equivp proofs failing)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2007 
diff
changeset
 | 
262  | 
|
| 
2125
 
60ee289a8c63
made out of STEPS a configuration value so that it can be set individually in each file
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2122 
diff
changeset
 | 
263  | 
setup STEPS_setup  | 
| 
 
60ee289a8c63
made out of STEPS a configuration value so that it can be set individually in each file
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2122 
diff
changeset
 | 
264  | 
|
| 1941 | 265  | 
ML {*
 | 
| 
2436
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2435 
diff
changeset
 | 
266  | 
fun nominal_datatype2 opt_thms_name dts bn_funs bn_eqs bclauses lthy =  | 
| 1941 | 267  | 
let  | 
| 
2294
 
72ad4e766acf
properly exported bn_descr
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2293 
diff
changeset
 | 
268  | 
(* definition of the raw datatypes *)  | 
| 2316 | 269  | 
val _ = warning "Definition of raw datatypes";  | 
| 
2337
 
b151399bd2c3
fixed according to changes in quotient
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2336 
diff
changeset
 | 
270  | 
val (raw_dt_names, raw_dts, raw_bclauses, raw_bn_funs, raw_bn_eqs, lthy0) =  | 
| 
2308
 
387fcbd33820
fixed problem with bn_info
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2306 
diff
changeset
 | 
271  | 
if get_STEPS lthy > 0  | 
| 
2410
 
2bbdb9c427b5
improved runtime slightly, by constructing an explicit size measure for the function definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2409 
diff
changeset
 | 
272  | 
then define_raw_dts dts bn_funs bn_eqs bclauses lthy  | 
| 
2046
 
73c50e913db6
tuned and added some comments to the code; added also an exception for early exit of the nominal2_cmd function
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2037 
diff
changeset
 | 
273  | 
else raise TEST lthy  | 
| 1941 | 274  | 
|
| 
2304
 
8a98171ba1fc
all raw definitions are defined using function
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2302 
diff
changeset
 | 
275  | 
val dtinfo = Datatype.the_info (ProofContext.theory_of lthy0) (hd raw_dt_names)  | 
| 
2143
 
871d8a5e0c67
somewhat simplified the main parsing function; failed to move a Note-statement to define_raw_perms
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2142 
diff
changeset
 | 
276  | 
  val {descr, sorts, ...} = dtinfo
 | 
| 2407 | 277  | 
|
| 
2400
 
c6d30d5f5ba1
defined qperms and qsizes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2399 
diff
changeset
 | 
278  | 
val raw_tys = all_dtyps descr sorts  | 
| 
 
c6d30d5f5ba1
defined qperms and qsizes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2399 
diff
changeset
 | 
279  | 
val raw_full_ty_names = map (fst o dest_Type) raw_tys  | 
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
280  | 
val tvs = hd raw_tys  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
281  | 
|> snd o dest_Type  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
282  | 
|> map dest_TFree  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
283  | 
|
| 
2400
 
c6d30d5f5ba1
defined qperms and qsizes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2399 
diff
changeset
 | 
284  | 
val dtinfos = map (Datatype.the_info (ProofContext.theory_of lthy0)) raw_full_ty_names  | 
| 
 
c6d30d5f5ba1
defined qperms and qsizes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2399 
diff
changeset
 | 
285  | 
|
| 2407 | 286  | 
val raw_cns_info = all_dtyp_constrs_types descr sorts  | 
287  | 
val raw_constrs = flat (map (map (fn (c, _, _, _) => c)) raw_cns_info)  | 
|
288  | 
||
| 
2400
 
c6d30d5f5ba1
defined qperms and qsizes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2399 
diff
changeset
 | 
289  | 
val raw_inject_thms = flat (map #inject dtinfos)  | 
| 
 
c6d30d5f5ba1
defined qperms and qsizes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2399 
diff
changeset
 | 
290  | 
val raw_distinct_thms = flat (map #distinct dtinfos)  | 
| 
 
c6d30d5f5ba1
defined qperms and qsizes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2399 
diff
changeset
 | 
291  | 
val raw_induct_thm = #induct dtinfo  | 
| 
 
c6d30d5f5ba1
defined qperms and qsizes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2399 
diff
changeset
 | 
292  | 
val raw_induct_thms = #inducts dtinfo  | 
| 
 
c6d30d5f5ba1
defined qperms and qsizes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2399 
diff
changeset
 | 
293  | 
val raw_exhaust_thms = map #exhaust dtinfos  | 
| 
 
c6d30d5f5ba1
defined qperms and qsizes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2399 
diff
changeset
 | 
294  | 
val raw_size_trms = map size_const raw_tys  | 
| 2388 | 295  | 
val raw_size_thms = Size.size_thms (ProofContext.theory_of lthy0) (hd raw_dt_names)  | 
296  | 
|> `(fn thms => (length thms) div 2)  | 
|
| 
2392
 
9294d7cec5e2
proved rsp-helper lemmas of size functions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2389 
diff
changeset
 | 
297  | 
|> uncurry drop  | 
| 2388 | 298  | 
|
| 
2409
 
83990a42a068
more tuning of the code
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2407 
diff
changeset
 | 
299  | 
(* definitions of raw permutations by primitive recursion *)  | 
| 2316 | 300  | 
val _ = warning "Definition of raw permutations";  | 
| 
2401
 
7645e18e8b19
modified the code for class instantiations (with help from Florian)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2400 
diff
changeset
 | 
301  | 
val ((raw_perm_funs, raw_perm_simps, raw_perm_laws), lthy2a) =  | 
| 
2308
 
387fcbd33820
fixed problem with bn_info
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2306 
diff
changeset
 | 
302  | 
if get_STEPS lthy0 > 1  | 
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
303  | 
then define_raw_perms raw_full_ty_names raw_tys tvs raw_constrs raw_induct_thm lthy0  | 
| 
2308
 
387fcbd33820
fixed problem with bn_info
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2306 
diff
changeset
 | 
304  | 
else raise TEST lthy0  | 
| 2144 | 305  | 
|
306  | 
(* noting the raw permutations as eqvt theorems *)  | 
|
| 2448 | 307  | 
val (_, lthy3) = Local_Theory.note ((Binding.empty, [eqvt_attr]), raw_perm_simps) lthy2a  | 
| 
2011
 
12ce87b55f97
tried to add some comments in the huge(!) nominal2_cmd function
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2008 
diff
changeset
 | 
308  | 
|
| 
2142
 
c39d4fe31100
moved the exporting part into the parser (this is still a hack); re-added CoreHaskell again to the examples - there seems to be a problem with the variable name pat
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2125 
diff
changeset
 | 
309  | 
(* definition of raw fv_functions *)  | 
| 2316 | 310  | 
val _ = warning "Definition of raw fv-functions";  | 
| 
2405
 
29ebbe56f450
also able to lift the bn_defs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2404 
diff
changeset
 | 
311  | 
val (raw_bns, raw_bn_defs, raw_bn_info, raw_bn_induct, lthy3a) =  | 
| 
2308
 
387fcbd33820
fixed problem with bn_info
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2306 
diff
changeset
 | 
312  | 
if get_STEPS lthy3 > 2  | 
| 
2410
 
2bbdb9c427b5
improved runtime slightly, by constructing an explicit size measure for the function definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2409 
diff
changeset
 | 
313  | 
then define_raw_bns raw_full_ty_names raw_dts raw_bn_funs raw_bn_eqs  | 
| 
 
2bbdb9c427b5
improved runtime slightly, by constructing an explicit size measure for the function definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2409 
diff
changeset
 | 
314  | 
(raw_inject_thms @ raw_distinct_thms) raw_size_thms lthy3  | 
| 
2046
 
73c50e913db6
tuned and added some comments to the code; added also an exception for early exit of the nominal2_cmd function
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2037 
diff
changeset
 | 
315  | 
else raise TEST lthy3  | 
| 2292 | 316  | 
|
| 
2308
 
387fcbd33820
fixed problem with bn_info
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2306 
diff
changeset
 | 
317  | 
val (raw_fvs, raw_fv_bns, raw_fv_defs, raw_fv_bns_induct, lthy3b) =  | 
| 
 
387fcbd33820
fixed problem with bn_info
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2306 
diff
changeset
 | 
318  | 
if get_STEPS lthy3a > 3  | 
| 2407 | 319  | 
then define_raw_fvs raw_full_ty_names raw_tys raw_cns_info raw_bn_info raw_bclauses  | 
| 
2410
 
2bbdb9c427b5
improved runtime slightly, by constructing an explicit size measure for the function definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2409 
diff
changeset
 | 
320  | 
(raw_inject_thms @ raw_distinct_thms) raw_size_thms lthy3a  | 
| 
2308
 
387fcbd33820
fixed problem with bn_info
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2306 
diff
changeset
 | 
321  | 
else raise TEST lthy3a  | 
| 
 
387fcbd33820
fixed problem with bn_info
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2306 
diff
changeset
 | 
322  | 
|
| 
2011
 
12ce87b55f97
tried to add some comments in the huge(!) nominal2_cmd function
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2008 
diff
changeset
 | 
323  | 
(* definition of raw alphas *)  | 
| 2316 | 324  | 
val _ = warning "Definition of alphas";  | 
| 
2300
 
9fb315392493
added FSet to the correct paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2299 
diff
changeset
 | 
325  | 
val (alpha_trms, alpha_bn_trms, alpha_intros, alpha_cases, alpha_induct, lthy4) =  | 
| 
2308
 
387fcbd33820
fixed problem with bn_info
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2306 
diff
changeset
 | 
326  | 
if get_STEPS lthy3b > 4  | 
| 2407 | 327  | 
then define_raw_alpha raw_full_ty_names raw_tys raw_cns_info raw_bn_info raw_bclauses raw_fvs lthy3b  | 
| 
2308
 
387fcbd33820
fixed problem with bn_info
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2306 
diff
changeset
 | 
328  | 
else raise TEST lthy3b  | 
| 
2336
 
f2d545b77b31
added definition of the quotient types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2324 
diff
changeset
 | 
329  | 
val alpha_tys = map (domain_type o fastype_of) alpha_trms  | 
| 
 
f2d545b77b31
added definition of the quotient types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2324 
diff
changeset
 | 
330  | 
|
| 
2300
 
9fb315392493
added FSet to the correct paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2299 
diff
changeset
 | 
331  | 
(* definition of alpha-distinct lemmas *)  | 
| 2316 | 332  | 
val _ = warning "Distinct theorems";  | 
| 2399 | 333  | 
val alpha_distincts =  | 
| 
2400
 
c6d30d5f5ba1
defined qperms and qsizes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2399 
diff
changeset
 | 
334  | 
mk_alpha_distincts lthy4 alpha_cases raw_distinct_thms alpha_trms raw_tys  | 
| 
2300
 
9fb315392493
added FSet to the correct paper
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2299 
diff
changeset
 | 
335  | 
|
| 2361 | 336  | 
(* definition of alpha_eq_iff lemmas *)  | 
| 
2438
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
337  | 
(* they have a funny shape for the simplifier ---- CHECK WHETHER NEEDED*)  | 
| 2316 | 338  | 
val _ = warning "Eq-iff theorems";  | 
| 
2387
 
082d9fd28f89
helper lemmas for rsp-lemmas
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2384 
diff
changeset
 | 
339  | 
val (alpha_eq_iff_simps, alpha_eq_iff) =  | 
| 
2295
 
8aff3f3ce47f
started to work on alpha
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2294 
diff
changeset
 | 
340  | 
if get_STEPS lthy > 5  | 
| 
2400
 
c6d30d5f5ba1
defined qperms and qsizes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2399 
diff
changeset
 | 
341  | 
then mk_alpha_eq_iff lthy4 alpha_intros raw_distinct_thms raw_inject_thms alpha_cases  | 
| 
2295
 
8aff3f3ce47f
started to work on alpha
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2294 
diff
changeset
 | 
342  | 
else raise TEST lthy4  | 
| 
2022
 
8ffede2c8ce9
Introduce eq_iff_simp to match the one from Parser.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2020 
diff
changeset
 | 
343  | 
|
| 2388 | 344  | 
(* proving equivariance lemmas for bns, fvs, size and alpha *)  | 
| 
2001
 
7c8242a02f39
NewParser with Parser functionality, but some cheats included since the order of datayupes is wrong.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
2000 
diff
changeset
 | 
345  | 
val _ = warning "Proving equivariance";  | 
| 
2406
 
428d9cb9a243
can also lift the various eqvt lemmas for bn, fv, fv_bn and size
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2405 
diff
changeset
 | 
346  | 
val raw_bn_eqvt =  | 
| 2298 | 347  | 
if get_STEPS lthy > 6  | 
| 
2405
 
29ebbe56f450
also able to lift the bn_defs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2404 
diff
changeset
 | 
348  | 
then raw_prove_eqvt raw_bns raw_bn_induct (raw_bn_defs @ raw_perm_simps) lthy4  | 
| 2298 | 349  | 
else raise TEST lthy4  | 
350  | 
||
| 
2406
 
428d9cb9a243
can also lift the various eqvt lemmas for bn, fv, fv_bn and size
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2405 
diff
changeset
 | 
351  | 
(* noting the raw_bn_eqvt lemmas in a temprorary theory *)  | 
| 2448 | 352  | 
val lthy_tmp = snd (Local_Theory.note ((Binding.empty, [eqvt_attr]), raw_bn_eqvt) lthy4)  | 
| 
2305
 
93ab397f5980
smaller code for raw-eqvt proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2304 
diff
changeset
 | 
353  | 
|
| 
2406
 
428d9cb9a243
can also lift the various eqvt lemmas for bn, fv, fv_bn and size
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2405 
diff
changeset
 | 
354  | 
val raw_fv_eqvt =  | 
| 2298 | 355  | 
if get_STEPS lthy > 7  | 
| 
2384
 
841b7e34e70a
fixed order of fold_union to make alpha and fv agree
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2378 
diff
changeset
 | 
356  | 
then raw_prove_eqvt (raw_fvs @ raw_fv_bns) raw_fv_bns_induct (raw_fv_defs @ raw_perm_simps)  | 
| 2388 | 357  | 
(Local_Theory.restore lthy_tmp)  | 
358  | 
else raise TEST lthy4  | 
|
359  | 
||
| 
2389
 
0f24c961b5f6
introduced a general alpha_prove method
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2388 
diff
changeset
 | 
360  | 
val raw_size_eqvt =  | 
| 2388 | 361  | 
if get_STEPS lthy > 8  | 
362  | 
then raw_prove_eqvt raw_size_trms raw_induct_thms (raw_size_thms @ raw_perm_simps)  | 
|
363  | 
(Local_Theory.restore lthy_tmp)  | 
|
| 
2389
 
0f24c961b5f6
introduced a general alpha_prove method
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2388 
diff
changeset
 | 
364  | 
      |> map (rewrite_rule @{thms permute_nat_def[THEN eq_reflection]})
 | 
| 
 
0f24c961b5f6
introduced a general alpha_prove method
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2388 
diff
changeset
 | 
365  | 
      |> map (fn thm => thm RS @{thm sym})
 | 
| 
2305
 
93ab397f5980
smaller code for raw-eqvt proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2304 
diff
changeset
 | 
366  | 
else raise TEST lthy4  | 
| 
 
93ab397f5980
smaller code for raw-eqvt proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2304 
diff
changeset
 | 
367  | 
|
| 2448 | 368  | 
val lthy5 = snd (Local_Theory.note ((Binding.empty, [eqvt_attr]), raw_fv_eqvt) lthy_tmp)  | 
| 2306 | 369  | 
|
| 
2336
 
f2d545b77b31
added definition of the quotient types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2324 
diff
changeset
 | 
370  | 
val (alpha_eqvt, lthy6) =  | 
| 2388 | 371  | 
if get_STEPS lthy > 9  | 
| 
2336
 
f2d545b77b31
added definition of the quotient types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2324 
diff
changeset
 | 
372  | 
then Nominal_Eqvt.equivariance true (alpha_trms @ alpha_bn_trms) alpha_induct alpha_intros lthy5  | 
| 
2311
 
4da5c5c29009
work on transitivity proof
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2309 
diff
changeset
 | 
373  | 
else raise TEST lthy4  | 
| 
 
4da5c5c29009
work on transitivity proof
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2309 
diff
changeset
 | 
374  | 
|
| 
 
4da5c5c29009
work on transitivity proof
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2309 
diff
changeset
 | 
375  | 
(* proving alpha equivalence *)  | 
| 
 
4da5c5c29009
work on transitivity proof
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2309 
diff
changeset
 | 
376  | 
val _ = warning "Proving equivalence"  | 
| 
 
4da5c5c29009
work on transitivity proof
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2309 
diff
changeset
 | 
377  | 
|
| 2316 | 378  | 
val alpha_refl_thms =  | 
| 2388 | 379  | 
if get_STEPS lthy > 10  | 
380  | 
then raw_prove_refl alpha_trms alpha_bn_trms alpha_intros raw_induct_thm lthy6  | 
|
| 
2336
 
f2d545b77b31
added definition of the quotient types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2324 
diff
changeset
 | 
381  | 
else raise TEST lthy6  | 
| 2316 | 382  | 
|
| 
2311
 
4da5c5c29009
work on transitivity proof
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2309 
diff
changeset
 | 
383  | 
val alpha_sym_thms =  | 
| 2388 | 384  | 
if get_STEPS lthy > 11  | 
| 
2336
 
f2d545b77b31
added definition of the quotient types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2324 
diff
changeset
 | 
385  | 
then raw_prove_sym (alpha_trms @ alpha_bn_trms) alpha_intros alpha_induct lthy6  | 
| 
 
f2d545b77b31
added definition of the quotient types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2324 
diff
changeset
 | 
386  | 
else raise TEST lthy6  | 
| 2298 | 387  | 
|
| 
2311
 
4da5c5c29009
work on transitivity proof
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2309 
diff
changeset
 | 
388  | 
val alpha_trans_thms =  | 
| 2388 | 389  | 
if get_STEPS lthy > 12  | 
| 
2400
 
c6d30d5f5ba1
defined qperms and qsizes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2399 
diff
changeset
 | 
390  | 
then raw_prove_trans (alpha_trms @ alpha_bn_trms) (raw_distinct_thms @ raw_inject_thms)  | 
| 
2336
 
f2d545b77b31
added definition of the quotient types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2324 
diff
changeset
 | 
391  | 
alpha_intros alpha_induct alpha_cases lthy6  | 
| 
 
f2d545b77b31
added definition of the quotient types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2324 
diff
changeset
 | 
392  | 
else raise TEST lthy6  | 
| 
2311
 
4da5c5c29009
work on transitivity proof
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2309 
diff
changeset
 | 
393  | 
|
| 
2404
 
66ae73fd66c0
added rsp-lemmas for alpha_bns
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2401 
diff
changeset
 | 
394  | 
val (alpha_equivp_thms, alpha_bn_equivp_thms) =  | 
| 2388 | 395  | 
if get_STEPS lthy > 13  | 
| 
2404
 
66ae73fd66c0
added rsp-lemmas for alpha_bns
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2401 
diff
changeset
 | 
396  | 
then raw_prove_equivp alpha_trms alpha_bn_trms alpha_refl_thms alpha_sym_thms  | 
| 
 
66ae73fd66c0
added rsp-lemmas for alpha_bns
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2401 
diff
changeset
 | 
397  | 
alpha_trans_thms lthy6  | 
| 
2336
 
f2d545b77b31
added definition of the quotient types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2324 
diff
changeset
 | 
398  | 
else raise TEST lthy6  | 
| 
2322
 
24de7e548094
proved eqvip theorems for alphas
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2321 
diff
changeset
 | 
399  | 
|
| 
2320
 
d835a2771608
prove that alpha implies alpha_bn (needed for rsp proofs)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2316 
diff
changeset
 | 
400  | 
(* proving alpha implies alpha_bn *)  | 
| 
 
d835a2771608
prove that alpha implies alpha_bn (needed for rsp proofs)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2316 
diff
changeset
 | 
401  | 
val _ = warning "Proving alpha implies bn"  | 
| 
 
d835a2771608
prove that alpha implies alpha_bn (needed for rsp proofs)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2316 
diff
changeset
 | 
402  | 
|
| 
 
d835a2771608
prove that alpha implies alpha_bn (needed for rsp proofs)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2316 
diff
changeset
 | 
403  | 
val alpha_bn_imp_thms =  | 
| 2388 | 404  | 
if get_STEPS lthy > 14  | 
| 
2336
 
f2d545b77b31
added definition of the quotient types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2324 
diff
changeset
 | 
405  | 
then raw_prove_bn_imp alpha_trms alpha_bn_trms alpha_intros alpha_induct lthy6  | 
| 
 
f2d545b77b31
added definition of the quotient types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2324 
diff
changeset
 | 
406  | 
else raise TEST lthy6  | 
| 
2322
 
24de7e548094
proved eqvip theorems for alphas
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2321 
diff
changeset
 | 
407  | 
|
| 
2397
 
c670a849af65
more experiments with lifting
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2396 
diff
changeset
 | 
408  | 
(* respectfulness proofs *)  | 
| 
2438
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
409  | 
val raw_funs_rsp_aux =  | 
| 
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
410  | 
if get_STEPS lthy > 15  | 
| 
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
411  | 
then raw_fv_bn_rsp_aux alpha_trms alpha_bn_trms raw_fvs  | 
| 
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
412  | 
raw_bns raw_fv_bns alpha_induct (raw_bn_defs @ raw_fv_defs) lthy6  | 
| 
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
413  | 
else raise TEST lthy6  | 
| 
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
414  | 
|
| 
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
415  | 
val raw_funs_rsp =  | 
| 
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
416  | 
if get_STEPS lthy > 16  | 
| 
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
417  | 
then map mk_funs_rsp raw_funs_rsp_aux  | 
| 
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
418  | 
else raise TEST lthy6  | 
| 2388 | 419  | 
|
| 
2438
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
420  | 
val raw_size_rsp =  | 
| 
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
421  | 
if get_STEPS lthy > 17  | 
| 
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
422  | 
then  | 
| 
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
423  | 
raw_size_rsp_aux (alpha_trms @ alpha_bn_trms) alpha_induct  | 
| 
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
424  | 
(raw_size_thms @ raw_size_eqvt) lthy6  | 
| 
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
425  | 
|> map mk_funs_rsp  | 
| 
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
426  | 
else raise TEST lthy6  | 
| 
2392
 
9294d7cec5e2
proved rsp-helper lemmas of size functions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2389 
diff
changeset
 | 
427  | 
|
| 
2438
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
428  | 
val raw_constrs_rsp =  | 
| 
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
429  | 
if get_STEPS lthy > 18  | 
| 
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
430  | 
then raw_constrs_rsp raw_constrs alpha_trms alpha_intros  | 
| 
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
431  | 
(alpha_bn_imp_thms @ raw_funs_rsp_aux) lthy6  | 
| 
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
432  | 
else raise TEST lthy6  | 
| 
2397
 
c670a849af65
more experiments with lifting
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2396 
diff
changeset
 | 
433  | 
|
| 
2438
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
434  | 
val alpha_permute_rsp =  | 
| 
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
435  | 
if get_STEPS lthy > 19  | 
| 
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
436  | 
then map mk_alpha_permute_rsp alpha_eqvt  | 
| 
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
437  | 
else raise TEST lthy6  | 
| 
2384
 
841b7e34e70a
fixed order of fold_union to make alpha and fv agree
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2378 
diff
changeset
 | 
438  | 
|
| 
2438
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
439  | 
val alpha_bn_rsp =  | 
| 
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
440  | 
if get_STEPS lthy > 20  | 
| 
2440
 
0a36825b16c1
"isabelle make test" makes all major examples....they work up to supp theorems (excluding)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2438 
diff
changeset
 | 
441  | 
then raw_alpha_bn_rsp alpha_bn_trms alpha_bn_equivp_thms alpha_bn_imp_thms  | 
| 
2438
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
442  | 
else raise TEST lthy6  | 
| 
2404
 
66ae73fd66c0
added rsp-lemmas for alpha_bns
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2401 
diff
changeset
 | 
443  | 
|
| 2398 | 444  | 
(* noting the quot_respects lemmas *)  | 
445  | 
val (_, lthy6a) =  | 
|
| 
2438
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
446  | 
if get_STEPS lthy > 21  | 
| 2448 | 447  | 
then Local_Theory.note ((Binding.empty, [rsp_attr]),  | 
| 
2404
 
66ae73fd66c0
added rsp-lemmas for alpha_bns
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2401 
diff
changeset
 | 
448  | 
raw_constrs_rsp @ raw_funs_rsp @ raw_size_rsp @ alpha_permute_rsp @ alpha_bn_rsp) lthy6  | 
| 2398 | 449  | 
else raise TEST lthy6  | 
450  | 
||
| 
2336
 
f2d545b77b31
added definition of the quotient types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2324 
diff
changeset
 | 
451  | 
(* defining the quotient type *)  | 
| 
 
f2d545b77b31
added definition of the quotient types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2324 
diff
changeset
 | 
452  | 
val _ = warning "Declaring the quotient types"  | 
| 
 
f2d545b77b31
added definition of the quotient types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2324 
diff
changeset
 | 
453  | 
val qty_descr = map (fn (vs, bind, mx, _) => (vs, bind, mx)) dts  | 
| 
2400
 
c6d30d5f5ba1
defined qperms and qsizes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2399 
diff
changeset
 | 
454  | 
|
| 
2336
 
f2d545b77b31
added definition of the quotient types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2324 
diff
changeset
 | 
455  | 
val (qty_infos, lthy7) =  | 
| 
2438
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
456  | 
if get_STEPS lthy > 22  | 
| 
2400
 
c6d30d5f5ba1
defined qperms and qsizes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2399 
diff
changeset
 | 
457  | 
then define_qtypes qty_descr alpha_tys alpha_trms alpha_equivp_thms lthy6a  | 
| 2398 | 458  | 
else raise TEST lthy6a  | 
| 
2336
 
f2d545b77b31
added definition of the quotient types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2324 
diff
changeset
 | 
459  | 
|
| 
 
f2d545b77b31
added definition of the quotient types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2324 
diff
changeset
 | 
460  | 
val qtys = map #qtyp qty_infos  | 
| 
2400
 
c6d30d5f5ba1
defined qperms and qsizes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2399 
diff
changeset
 | 
461  | 
val qty_full_names = map (fst o dest_Type) qtys  | 
| 
 
c6d30d5f5ba1
defined qperms and qsizes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2399 
diff
changeset
 | 
462  | 
val qty_names = map Long_Name.base_name qty_full_names  | 
| 
 
c6d30d5f5ba1
defined qperms and qsizes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2399 
diff
changeset
 | 
463  | 
|
| 
2339
 
1e0b3992189c
more quotient-definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2338 
diff
changeset
 | 
464  | 
(* defining of quotient term-constructors, binding functions, free vars functions *)  | 
| 
2378
 
2f13fe48c877
updated to new Isabelle; made FSet more "quiet"
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2361 
diff
changeset
 | 
465  | 
val _ = warning "Defining the quotient constants"  | 
| 2346 | 466  | 
val qconstrs_descr =  | 
| 2338 | 467  | 
flat (map (fn (_, _, _, cs) => map (fn (b, _, mx) => (Name.of_binding b, mx)) cs) dts)  | 
| 
2400
 
c6d30d5f5ba1
defined qperms and qsizes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2399 
diff
changeset
 | 
468  | 
|> map2 (fn t => fn (b, mx) => (b, t, mx)) raw_constrs  | 
| 2338 | 469  | 
|
| 
2339
 
1e0b3992189c
more quotient-definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2338 
diff
changeset
 | 
470  | 
val qbns_descr =  | 
| 2346 | 471  | 
map2 (fn (b, _, mx) => fn t => (Name.of_binding b, t, mx)) bn_funs raw_bns  | 
| 
2339
 
1e0b3992189c
more quotient-definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2338 
diff
changeset
 | 
472  | 
|
| 
 
1e0b3992189c
more quotient-definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2338 
diff
changeset
 | 
473  | 
val qfvs_descr =  | 
| 2346 | 474  | 
    map2 (fn n => fn t => ("fv_" ^ n, t, NoSyn)) qty_names raw_fvs
 | 
| 
2339
 
1e0b3992189c
more quotient-definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2338 
diff
changeset
 | 
475  | 
|
| 2346 | 476  | 
val qfv_bns_descr =  | 
| 2398 | 477  | 
    map2 (fn (b, _, _) => fn t => ("fv_" ^ Name.of_binding b, t, NoSyn)) bn_funs raw_fv_bns
 | 
| 
2339
 
1e0b3992189c
more quotient-definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2338 
diff
changeset
 | 
478  | 
|
| 
2384
 
841b7e34e70a
fixed order of fold_union to make alpha and fv agree
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2378 
diff
changeset
 | 
479  | 
val qalpha_bns_descr =  | 
| 
 
841b7e34e70a
fixed order of fold_union to make alpha and fv agree
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2378 
diff
changeset
 | 
480  | 
    map2 (fn (b, _, _) => fn t => ("alpha_" ^ Name.of_binding b, t, NoSyn)) bn_funs  alpha_bn_trms
 | 
| 
 
841b7e34e70a
fixed order of fold_union to make alpha and fv agree
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2378 
diff
changeset
 | 
481  | 
|
| 2398 | 482  | 
val qperm_descr =  | 
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
483  | 
    map2 (fn n => fn t => ("permute_" ^ n, Type.legacy_freeze t, NoSyn)) qty_names raw_perm_funs
 | 
| 2398 | 484  | 
|
| 
2400
 
c6d30d5f5ba1
defined qperms and qsizes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2399 
diff
changeset
 | 
485  | 
val qsize_descr =  | 
| 
 
c6d30d5f5ba1
defined qperms and qsizes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2399 
diff
changeset
 | 
486  | 
    map2 (fn n => fn t => ("size_" ^ n, t, NoSyn)) qty_names raw_size_trms
 | 
| 
 
c6d30d5f5ba1
defined qperms and qsizes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2399 
diff
changeset
 | 
487  | 
|
| 
2384
 
841b7e34e70a
fixed order of fold_union to make alpha and fv agree
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2378 
diff
changeset
 | 
488  | 
val (((((qconstrs_info, qbns_info), qfvs_info), qfv_bns_info), qalpha_bns_info), lthy8) =  | 
| 
2438
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
489  | 
if get_STEPS lthy > 23  | 
| 2346 | 490  | 
then  | 
491  | 
lthy7  | 
|
| 
2400
 
c6d30d5f5ba1
defined qperms and qsizes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2399 
diff
changeset
 | 
492  | 
|> define_qconsts qtys qconstrs_descr  | 
| 
 
c6d30d5f5ba1
defined qperms and qsizes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2399 
diff
changeset
 | 
493  | 
||>> define_qconsts qtys qbns_descr  | 
| 
 
c6d30d5f5ba1
defined qperms and qsizes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2399 
diff
changeset
 | 
494  | 
||>> define_qconsts qtys qfvs_descr  | 
| 
 
c6d30d5f5ba1
defined qperms and qsizes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2399 
diff
changeset
 | 
495  | 
||>> define_qconsts qtys qfv_bns_descr  | 
| 
 
c6d30d5f5ba1
defined qperms and qsizes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2399 
diff
changeset
 | 
496  | 
||>> define_qconsts qtys qalpha_bns_descr  | 
| 2338 | 497  | 
else raise TEST lthy7  | 
498  | 
||
| 
2400
 
c6d30d5f5ba1
defined qperms and qsizes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2399 
diff
changeset
 | 
499  | 
(* definition of the quotient permfunctions and pt-class *)  | 
| 
 
c6d30d5f5ba1
defined qperms and qsizes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2399 
diff
changeset
 | 
500  | 
val lthy9 =  | 
| 
2438
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
501  | 
if get_STEPS lthy > 24  | 
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
502  | 
then define_qperms qtys qty_full_names tvs qperm_descr raw_perm_laws lthy8  | 
| 
2400
 
c6d30d5f5ba1
defined qperms and qsizes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2399 
diff
changeset
 | 
503  | 
else raise TEST lthy8  | 
| 
 
c6d30d5f5ba1
defined qperms and qsizes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2399 
diff
changeset
 | 
504  | 
|
| 
2401
 
7645e18e8b19
modified the code for class instantiations (with help from Florian)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2400 
diff
changeset
 | 
505  | 
val lthy9a =  | 
| 
2438
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
506  | 
if get_STEPS lthy > 25  | 
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
507  | 
then define_qsizes qtys qty_full_names tvs qsize_descr lthy9  | 
| 
2400
 
c6d30d5f5ba1
defined qperms and qsizes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2399 
diff
changeset
 | 
508  | 
else raise TEST lthy9  | 
| 2398 | 509  | 
|
| 2448 | 510  | 
val qtrms = map #qconst qconstrs_info  | 
| 2346 | 511  | 
val qbns = map #qconst qbns_info  | 
512  | 
val qfvs = map #qconst qfvs_info  | 
|
513  | 
val qfv_bns = map #qconst qfv_bns_info  | 
|
| 
2384
 
841b7e34e70a
fixed order of fold_union to make alpha and fv agree
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2378 
diff
changeset
 | 
514  | 
val qalpha_bns = map #qconst qalpha_bns_info  | 
| 2434 | 515  | 
|
516  | 
(* lifting of the theorems *)  | 
|
517  | 
val _ = warning "Lifting of Theorems"  | 
|
518  | 
||
519  | 
  val eq_iff_simps = @{thms alphas permute_prod.simps prod_fv.simps prod_alpha_def prod_rel.simps
 | 
|
520  | 
prod.cases}  | 
|
521  | 
||
522  | 
val ((((((qdistincts, qeq_iffs), qfv_defs), qbn_defs), qperm_simps), qfv_qbn_eqvts), lthyA) =  | 
|
| 
2438
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
523  | 
if get_STEPS lthy > 26  | 
| 2434 | 524  | 
then  | 
525  | 
lthy9a  | 
|
526  | 
|> lift_thms qtys [] alpha_distincts  | 
|
527  | 
||>> lift_thms qtys eq_iff_simps alpha_eq_iff  | 
|
528  | 
||>> lift_thms qtys [] raw_fv_defs  | 
|
529  | 
||>> lift_thms qtys [] raw_bn_defs  | 
|
530  | 
||>> lift_thms qtys [] raw_perm_simps  | 
|
531  | 
||>> lift_thms qtys [] (raw_fv_eqvt @ raw_bn_eqvt)  | 
|
532  | 
else raise TEST lthy9a  | 
|
533  | 
||
534  | 
val (((qsize_eqvt, [qinduct]), qexhausts), lthyB) =  | 
|
| 
2438
 
abafea9b39bb
corrected bug with fv-function generation (that was the problem with recursive binders)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2436 
diff
changeset
 | 
535  | 
if get_STEPS lthy > 27  | 
| 2434 | 536  | 
then  | 
537  | 
lthyA  | 
|
538  | 
|> lift_thms qtys [] raw_size_eqvt  | 
|
539  | 
||>> lift_thms qtys [] [raw_induct_thm]  | 
|
540  | 
||>> lift_thms qtys [] raw_exhaust_thms  | 
|
541  | 
else raise TEST lthyA  | 
|
542  | 
||
| 
2451
 
d2e929f51fa9
added fs-instance proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2450 
diff
changeset
 | 
543  | 
(* supports lemmas *)  | 
| 2448 | 544  | 
val qsupports_thms =  | 
545  | 
if get_STEPS lthy > 28  | 
|
546  | 
then prove_supports lthyB qperm_simps qtrms  | 
|
547  | 
else raise TEST lthyB  | 
|
548  | 
||
| 
2451
 
d2e929f51fa9
added fs-instance proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2450 
diff
changeset
 | 
549  | 
(* finite supp lemmas *)  | 
| 
2450
 
217ef3e4282e
added proofs for fsupp properties
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2448 
diff
changeset
 | 
550  | 
val qfsupp_thms =  | 
| 
 
217ef3e4282e
added proofs for fsupp properties
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2448 
diff
changeset
 | 
551  | 
if get_STEPS lthy > 29  | 
| 
 
217ef3e4282e
added proofs for fsupp properties
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2448 
diff
changeset
 | 
552  | 
then prove_fsupp lthyB qtys qinduct qsupports_thms  | 
| 
 
217ef3e4282e
added proofs for fsupp properties
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2448 
diff
changeset
 | 
553  | 
else raise TEST lthyB  | 
| 
 
217ef3e4282e
added proofs for fsupp properties
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2448 
diff
changeset
 | 
554  | 
|
| 
2451
 
d2e929f51fa9
added fs-instance proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2450 
diff
changeset
 | 
555  | 
(* fs instances *)  | 
| 
 
d2e929f51fa9
added fs-instance proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2450 
diff
changeset
 | 
556  | 
val lthyC =  | 
| 
 
d2e929f51fa9
added fs-instance proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2450 
diff
changeset
 | 
557  | 
if get_STEPS lthy > 30  | 
| 
 
d2e929f51fa9
added fs-instance proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2450 
diff
changeset
 | 
558  | 
then fs_instance qtys qty_full_names tvs qfsupp_thms lthyB  | 
| 
 
d2e929f51fa9
added fs-instance proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2450 
diff
changeset
 | 
559  | 
else raise TEST lthyB  | 
| 2448 | 560  | 
|
| 
2436
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2435 
diff
changeset
 | 
561  | 
(* noting the theorems *)  | 
| 
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2435 
diff
changeset
 | 
562  | 
|
| 
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2435 
diff
changeset
 | 
563  | 
(* generating the prefix for the theorem names *)  | 
| 
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2435 
diff
changeset
 | 
564  | 
val thms_name =  | 
| 
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2435 
diff
changeset
 | 
565  | 
the_default (Binding.name (space_implode "_" qty_names)) opt_thms_name  | 
| 
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2435 
diff
changeset
 | 
566  | 
fun thms_suffix s = Binding.qualified true s thms_name  | 
| 
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2435 
diff
changeset
 | 
567  | 
|
| 
2451
 
d2e929f51fa9
added fs-instance proofs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2450 
diff
changeset
 | 
568  | 
val (_, lthy9') = lthyC  | 
| 
2436
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2435 
diff
changeset
 | 
569  | 
|> Local_Theory.note ((thms_suffix "distinct", []), qdistincts)  | 
| 
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2435 
diff
changeset
 | 
570  | 
||>> Local_Theory.note ((thms_suffix "eq_iff", []), qeq_iffs)  | 
| 
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2435 
diff
changeset
 | 
571  | 
||>> Local_Theory.note ((thms_suffix "fv_defs", []), qfv_defs)  | 
| 
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2435 
diff
changeset
 | 
572  | 
||>> Local_Theory.note ((thms_suffix "bn_defs", []), qbn_defs)  | 
| 2448 | 573  | 
||>> Local_Theory.note ((thms_suffix "perm_simps", [eqvt_attr, simp_attr]), qperm_simps)  | 
| 
2436
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2435 
diff
changeset
 | 
574  | 
||>> Local_Theory.note ((thms_suffix "fv_bn_eqvt", []), qfv_qbn_eqvts)  | 
| 
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2435 
diff
changeset
 | 
575  | 
||>> Local_Theory.note ((thms_suffix "size_eqvt", []), qsize_eqvt)  | 
| 
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2435 
diff
changeset
 | 
576  | 
||>> Local_Theory.note ((thms_suffix "induct", []), [qinduct])  | 
| 
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2435 
diff
changeset
 | 
577  | 
||>> Local_Theory.note ((thms_suffix "exhaust", []), qexhausts)  | 
| 2448 | 578  | 
||>> Local_Theory.note ((thms_suffix "supports", []), qsupports_thms)  | 
| 
2450
 
217ef3e4282e
added proofs for fsupp properties
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2448 
diff
changeset
 | 
579  | 
||>> Local_Theory.note ((thms_suffix "fsupp", []), qfsupp_thms)  | 
| 1941 | 580  | 
in  | 
| 
2435
 
3772bb3bd7ce
cleaning of unused files and code
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2434 
diff
changeset
 | 
581  | 
(0, lthy9')  | 
| 
2046
 
73c50e913db6
tuned and added some comments to the code; added also an exception for early exit of the nominal2_cmd function
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2037 
diff
changeset
 | 
582  | 
end handle TEST ctxt => (0, ctxt)  | 
| 1941 | 583  | 
*}  | 
584  | 
||
585  | 
section {* Preparing and parsing of the specification *}
 | 
|
586  | 
||
587  | 
ML {* 
 | 
|
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
588  | 
(* generates the parsed datatypes and  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
589  | 
declares the constructors  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
590  | 
*)  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
591  | 
fun prepare_dts dt_strs thy =  | 
| 1941 | 592  | 
let  | 
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
593  | 
fun inter_fs_sort thy (a, S) =  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
594  | 
    (a, Type.inter_sort (Sign.tsig_of thy) (@{sort fs}, S)) 
 | 
| 1941 | 595  | 
|
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
596  | 
fun mk_type tname sorts (cname, cargs, mx) =  | 
| 1941 | 597  | 
let  | 
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
598  | 
val full_tname = Sign.full_name thy tname  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
599  | 
val ty = Type (full_tname, map (TFree o inter_fs_sort thy) sorts)  | 
| 1941 | 600  | 
in  | 
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
601  | 
(cname, cargs ---> ty, mx)  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
602  | 
end  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
603  | 
|
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
604  | 
fun prep_constr (cname, cargs, mx, _) (constrs, sorts) =  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
605  | 
let  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
606  | 
val (cargs', sorts') =  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
607  | 
fold_map (Datatype.read_typ thy) (map snd cargs) sorts  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
608  | 
|>> map (map_type_tfree (TFree o inter_fs_sort thy))  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
609  | 
in  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
610  | 
(constrs @ [(cname, cargs', mx)], sorts')  | 
| 1941 | 611  | 
end  | 
612  | 
||
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
613  | 
fun prep_dts (tvs, tname, mx, constrs) (constr_trms, dts, sorts) =  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
614  | 
let  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
615  | 
val (constrs', sorts') =  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
616  | 
fold prep_constr constrs ([], sorts)  | 
| 1941 | 617  | 
|
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
618  | 
val constr_trms' =  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
619  | 
map (mk_type tname (rev sorts')) constrs'  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
620  | 
in  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
621  | 
(constr_trms @ constr_trms', dts @ [(tvs, tname, mx, constrs')], sorts')  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
622  | 
end  | 
| 
2425
 
715ab84065a0
nominal_datatypes with type variables do not work
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2424 
diff
changeset
 | 
623  | 
|
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
624  | 
val (constr_trms, dts, sorts) = fold prep_dts dt_strs ([], [], []);  | 
| 1941 | 625  | 
in  | 
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
626  | 
thy  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
627  | 
|> Sign.add_consts_i constr_trms  | 
| 1941 | 628  | 
|> pair dts  | 
629  | 
end  | 
|
630  | 
*}  | 
|
631  | 
||
632  | 
ML {*
 | 
|
633  | 
(* parsing the binding function specification and *)  | 
|
634  | 
(* declaring the functions in the local theory *)  | 
|
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
635  | 
fun prepare_bn_funs bn_fun_strs bn_eq_strs thy =  | 
| 1941 | 636  | 
let  | 
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
637  | 
val lthy = Named_Target.theory_init thy  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
638  | 
|
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
639  | 
val ((bn_funs, bn_eqs), lthy') =  | 
| 1941 | 640  | 
Specification.read_spec bn_fun_strs bn_eq_strs lthy  | 
641  | 
||
642  | 
fun prep_bn_fun ((bn, T), mx) = (bn, T, mx)  | 
|
643  | 
||
644  | 
val bn_funs' = map prep_bn_fun bn_funs  | 
|
645  | 
in  | 
|
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
646  | 
(Local_Theory.exit_global lthy')  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
647  | 
|> Sign.add_consts_i bn_funs'  | 
| 1941 | 648  | 
|> pair (bn_funs', bn_eqs)  | 
649  | 
end  | 
|
650  | 
*}  | 
|
651  | 
||
652  | 
text {* associates every SOME with the index in the list; drops NONEs *}
 | 
|
653  | 
ML {*
 | 
|
654  | 
fun indexify xs =  | 
|
655  | 
let  | 
|
656  | 
fun mapp _ [] = []  | 
|
657  | 
| mapp i (NONE :: xs) = mapp (i + 1) xs  | 
|
658  | 
| mapp i (SOME x :: xs) = (x, i) :: mapp (i + 1) xs  | 
|
659  | 
in  | 
|
660  | 
mapp 0 xs  | 
|
661  | 
end  | 
|
662  | 
||
663  | 
fun index_lookup xs x =  | 
|
664  | 
case AList.lookup (op=) xs x of  | 
|
665  | 
SOME x => x  | 
|
666  | 
  | NONE => error ("Cannot find " ^ x ^ " as argument annotation.");
 | 
|
667  | 
*}  | 
|
668  | 
||
669  | 
ML {*
 | 
|
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
670  | 
fun prepare_bclauses dt_strs thy =  | 
| 1941 | 671  | 
let  | 
672  | 
val annos_bclauses =  | 
|
673  | 
get_cnstrs dt_strs  | 
|
674  | 
|> map (map (fn (_, antys, _, bns) => (map fst antys, bns)))  | 
|
675  | 
||
676  | 
fun prep_binder env bn_str =  | 
|
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
677  | 
case (Syntax.read_term_global thy bn_str) of  | 
| 1941 | 678  | 
Free (x, _) => (NONE, index_lookup env x)  | 
679  | 
| Const (a, T) $ Free (x, _) => (SOME (Const (a, T)), index_lookup env x)  | 
|
680  | 
    | _ => error ("The term " ^ bn_str ^ " is not allowed as binding function.")
 | 
|
681  | 
||
682  | 
fun prep_body env bn_str = index_lookup env bn_str  | 
|
683  | 
||
684  | 
fun prep_bclause env (mode, binders, bodies) =  | 
|
685  | 
let  | 
|
686  | 
val binders' = map (prep_binder env) binders  | 
|
687  | 
val bodies' = map (prep_body env) bodies  | 
|
688  | 
in  | 
|
| 
2424
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2411 
diff
changeset
 | 
689  | 
BC (mode, binders', bodies')  | 
| 1941 | 690  | 
end  | 
691  | 
||
692  | 
fun prep_bclauses (annos, bclause_strs) =  | 
|
693  | 
let  | 
|
694  | 
val env = indexify annos (* for every label, associate the index *)  | 
|
695  | 
in  | 
|
696  | 
map (prep_bclause env) bclause_strs  | 
|
697  | 
end  | 
|
698  | 
in  | 
|
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
699  | 
(map (map prep_bclauses) annos_bclauses, thy)  | 
| 1941 | 700  | 
end  | 
701  | 
*}  | 
|
702  | 
||
| 1943 | 703  | 
text {* 
 | 
704  | 
adds an empty binding clause for every argument  | 
|
705  | 
that is not already part of a binding clause  | 
|
706  | 
*}  | 
|
707  | 
||
| 1941 | 708  | 
ML {*
 | 
709  | 
fun included i bcs =  | 
|
710  | 
let  | 
|
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
711  | 
fun incl (BC (_, bns, bds)) =  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
712  | 
member (op =) (map snd bns) i orelse member (op =) bds i  | 
| 1941 | 713  | 
in  | 
714  | 
exists incl bcs  | 
|
715  | 
end  | 
|
716  | 
*}  | 
|
717  | 
||
718  | 
ML {* 
 | 
|
719  | 
fun complete dt_strs bclauses =  | 
|
720  | 
let  | 
|
721  | 
val args =  | 
|
722  | 
get_cnstrs dt_strs  | 
|
723  | 
|> map (map (fn (_, antys, _, _) => length antys))  | 
|
724  | 
||
725  | 
fun complt n bcs =  | 
|
726  | 
let  | 
|
| 
2288
 
3b83960f9544
new fv/fv_bn function (supp breaks now); exported raw perms and raw funs into separate ML-files
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2146 
diff
changeset
 | 
727  | 
fun add bcs i = (if included i bcs then [] else [BC (Lst, [], [i])])  | 
| 1941 | 728  | 
in  | 
729  | 
bcs @ (flat (map_range (add bcs) n))  | 
|
730  | 
end  | 
|
731  | 
in  | 
|
732  | 
map2 (map2 complt) args bclauses  | 
|
733  | 
end  | 
|
734  | 
*}  | 
|
735  | 
||
736  | 
ML {*
 | 
|
| 
2436
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2435 
diff
changeset
 | 
737  | 
fun nominal_datatype2_cmd (opt_thms_name, dt_strs, bn_fun_strs, bn_eq_strs) lthy =  | 
| 1941 | 738  | 
let  | 
| 
2436
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2435 
diff
changeset
 | 
739  | 
val pre_typs =  | 
| 
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2435 
diff
changeset
 | 
740  | 
map (fn (tvs, tname, mx, _) => (tname, length tvs, mx)) dt_strs  | 
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
741  | 
|
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
742  | 
(* this theory is used just for parsing *)  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
743  | 
val thy = ProofContext.theory_of lthy  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
744  | 
val tmp_thy = Theory.copy thy  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
745  | 
|
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
746  | 
val (((dts, (bn_funs, bn_eqs)), bclauses), tmp_thy') =  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
747  | 
tmp_thy  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
748  | 
|> Sign.add_types pre_typs  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
749  | 
|> prepare_dts dt_strs  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
750  | 
||>> prepare_bn_funs bn_fun_strs bn_eq_strs  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
751  | 
||>> prepare_bclauses dt_strs  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
752  | 
|
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
753  | 
val bclauses' = complete dt_strs bclauses  | 
| 1941 | 754  | 
in  | 
| 
2436
 
3885dc2669f9
cleaned up (almost completely) the examples
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2435 
diff
changeset
 | 
755  | 
timeit (fn () => nominal_datatype2 opt_thms_name dts bn_funs bn_eqs bclauses' lthy |> snd)  | 
| 1941 | 756  | 
end  | 
| 
2424
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2411 
diff
changeset
 | 
757  | 
*}  | 
| 1941 | 758  | 
|
| 
2424
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2411 
diff
changeset
 | 
759  | 
ML {* 
 | 
| 
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2411 
diff
changeset
 | 
760  | 
(* nominal datatype parser *)  | 
| 
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2411 
diff
changeset
 | 
761  | 
local  | 
| 
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2411 
diff
changeset
 | 
762  | 
structure P = Parse;  | 
| 
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2411 
diff
changeset
 | 
763  | 
structure S = Scan  | 
| 
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2411 
diff
changeset
 | 
764  | 
|
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
765  | 
fun triple ((x, y), z) = (x, y, z)  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
766  | 
fun tuple1 ((x, y, z), u) = (x, y, z, u)  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
767  | 
fun tuple2 (((x, y), z), u) = (x, y, u, z)  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
768  | 
fun tuple3 ((x, y), (z, u)) = (x, y, z, u)  | 
| 
2424
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2411 
diff
changeset
 | 
769  | 
in  | 
| 
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2411 
diff
changeset
 | 
770  | 
|
| 
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2411 
diff
changeset
 | 
771  | 
val _ = Keyword.keyword "bind"  | 
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
772  | 
|
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
773  | 
val opt_name = Scan.option (P.binding --| Args.colon)  | 
| 
2424
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2411 
diff
changeset
 | 
774  | 
|
| 
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2411 
diff
changeset
 | 
775  | 
val anno_typ = S.option (P.name --| P.$$$ "::") -- P.typ  | 
| 
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2411 
diff
changeset
 | 
776  | 
|
| 
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2411 
diff
changeset
 | 
777  | 
val bind_mode = P.$$$ "bind" |--  | 
| 
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2411 
diff
changeset
 | 
778  | 
S.optional (Args.parens  | 
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
779  | 
(Args.$$$ "list" >> K Lst || Args.$$$ "set" >> K Set || Args.$$$ "res" >> K Res)) Lst  | 
| 
2424
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2411 
diff
changeset
 | 
780  | 
|
| 
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2411 
diff
changeset
 | 
781  | 
val bind_clauses =  | 
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
782  | 
P.enum "," (bind_mode -- S.repeat1 P.term -- (P.$$$ "in" |-- S.repeat1 P.name) >> triple)  | 
| 
2424
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2411 
diff
changeset
 | 
783  | 
|
| 
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2411 
diff
changeset
 | 
784  | 
val cnstr_parser =  | 
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
785  | 
P.binding -- S.repeat anno_typ -- bind_clauses -- P.opt_mixfix >> tuple2  | 
| 
2424
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2411 
diff
changeset
 | 
786  | 
|
| 
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2411 
diff
changeset
 | 
787  | 
(* datatype parser *)  | 
| 
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2411 
diff
changeset
 | 
788  | 
val dt_parser =  | 
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
789  | 
(P.type_args -- P.binding -- P.opt_mixfix >> triple) --  | 
| 
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
790  | 
(P.$$$ "=" |-- P.enum1 "|" cnstr_parser) >> tuple1  | 
| 
2424
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2411 
diff
changeset
 | 
791  | 
|
| 
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2411 
diff
changeset
 | 
792  | 
(* binding function parser *)  | 
| 
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2411 
diff
changeset
 | 
793  | 
val bnfun_parser =  | 
| 
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2411 
diff
changeset
 | 
794  | 
S.optional (P.$$$ "binder" |-- P.fixes -- Parse_Spec.where_alt_specs) ([], [])  | 
| 
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2411 
diff
changeset
 | 
795  | 
|
| 
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2411 
diff
changeset
 | 
796  | 
(* main parser *)  | 
| 
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2411 
diff
changeset
 | 
797  | 
val main_parser =  | 
| 
2431
 
331873ebc5cd
can now deal with type variables in nominal datatype definitions
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2430 
diff
changeset
 | 
798  | 
opt_name -- P.and_list1 dt_parser -- bnfun_parser >> tuple3  | 
| 
2424
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2411 
diff
changeset
 | 
799  | 
|
| 
 
621ebd8b13c4
changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
 
Christian Urban <urbanc@in.tum.de> 
parents: 
2411 
diff
changeset
 | 
800  | 
end  | 
| 1941 | 801  | 
|
802  | 
(* Command Keyword *)  | 
|
| 2168 | 803  | 
val _ = Outer_Syntax.local_theory "nominal_datatype" "test" Keyword.thy_decl  | 
| 1941 | 804  | 
(main_parser >> nominal_datatype2_cmd)  | 
805  | 
*}  | 
|
806  | 
||
| 2292 | 807  | 
|
| 1941 | 808  | 
end  | 
809  | 
||
810  | 
||
811  |