author | Cezary Kaliszyk <kaliszyk@in.tum.de> |
Wed, 27 Jan 2010 17:18:30 +0100 | |
changeset 970 | c16135580a45 |
parent 969 | cc5d18c78446 |
child 971 | af57c9723fae |
permissions | -rw-r--r-- |
954
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1 |
theory Test |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
2 |
imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
3 |
begin |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
4 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
5 |
atom_decl name |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
6 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
7 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
8 |
(* test for how to add datatypes *) |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
9 |
setup {* |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
10 |
Datatype.datatype_cmd |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
11 |
["foo"] |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
12 |
[([], @{binding "foo"}, NoSyn, |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
13 |
[(@{binding "zero"}, [], NoSyn),(@{binding "suc"}, ["foo"], NoSyn)]) |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
14 |
] |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
15 |
*} |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
16 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
17 |
thm foo.recs |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
18 |
thm foo.induct |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
19 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
20 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
21 |
ML{* |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
22 |
fun filtered_input str = |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
23 |
filter OuterLex.is_proper (OuterSyntax.scan Position.none str) |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
24 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
25 |
fun parse p input = Scan.finite OuterLex.stopper (Scan.error p) input |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
26 |
*} |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
27 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
28 |
(* type plus possible annotation *) |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
29 |
ML {* |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
30 |
val anno_typ = |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
31 |
(Scan.option (OuterParse.name --| OuterParse.$$$ "::")) -- OuterParse.typ |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
32 |
*} |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
33 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
34 |
ML {* |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
35 |
parse (Scan.repeat anno_typ) (filtered_input "x::string bool") |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
36 |
*} |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
37 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
38 |
ML {* OuterParse.enum "," ; OuterParse.and_list1 *} |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
39 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
40 |
(* binding specification *) |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
41 |
ML {* |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
42 |
datatype bind = B of string * string | BS of string * string |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
43 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
44 |
val _ = OuterKeyword.keyword "bind" |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
45 |
val _ = OuterKeyword.keyword "bindset" |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
46 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
47 |
val bind_parse_aux = |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
48 |
(OuterParse.$$$ "bind" >> (K B)) |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
49 |
|| (OuterParse.$$$ "bindset" >> (K BS)) |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
50 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
51 |
(* should use and_list *) |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
52 |
val bind_parser = |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
53 |
OuterParse.enum "," |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
54 |
((bind_parse_aux -- OuterParse.term) -- |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
55 |
(OuterParse.$$$ "in" |-- OuterParse.name) >> (fn ((bind, s1), s2) => bind (s1,s2))) |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
56 |
*} |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
57 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
58 |
(* datatype parser *) |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
59 |
ML {* |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
60 |
val dt_parser = |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
61 |
OuterParse.type_args -- OuterParse.binding -- OuterParse.opt_infix -- |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
62 |
(OuterParse.$$$ "=" |-- OuterParse.enum1 "|" |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
63 |
(OuterParse.binding -- (Scan.repeat anno_typ) -- bind_parser -- OuterParse.opt_mixfix)) |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
64 |
*} |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
65 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
66 |
(* function equation parser *) |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
67 |
ML {* |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
68 |
val fun_parser = |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
69 |
Scan.optional |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
70 |
((OuterParse.$$$ "binder") |-- OuterParse.fixes -- SpecParse.where_alt_specs) ([],[]) |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
71 |
*} |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
72 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
73 |
(* main parser *) |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
74 |
ML {* |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
75 |
val parser = |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
76 |
OuterParse.and_list1 dt_parser -- fun_parser |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
77 |
*} |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
78 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
79 |
(* printing stuff *) |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
80 |
ML {* |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
81 |
fun print_str_option NONE = "NONE" |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
82 |
| print_str_option (SOME s) = (s:bstring) |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
83 |
*} |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
84 |
|
969
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
85 |
ML {* |
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
86 |
fun print_anno_typ lthy (NONE, ty) = Syntax.string_of_typ lthy ty |
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
87 |
| print_anno_typ lthy (SOME x, ty) = x ^ ":" ^ Syntax.string_of_typ lthy ty |
954
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
88 |
*} |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
89 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
90 |
ML {* |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
91 |
fun print_bind (B (s1, s2)) = "bind " ^ s1 ^ " in " ^ s2 |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
92 |
| print_bind (BS (s1, s2)) = "bindset " ^ s1 ^ " in " ^ s2 |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
93 |
*} |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
94 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
95 |
ML {* |
969
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
96 |
fun print_constr lthy (((name, anno_tys), bds), syn) = |
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
97 |
(Binding.name_of name ^ " of " ^ (commas (map (print_anno_typ lthy) anno_tys)) ^ " " |
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
98 |
^ (commas (map print_bind bds)) ^ " " |
954
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
99 |
^ (Pretty.str_of (Syntax.pretty_mixfix syn))) |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
100 |
*} |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
101 |
|
966
8ba35ce16f7e
Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
962
diff
changeset
|
102 |
ML Local_Theory.exit_global |
8ba35ce16f7e
Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
962
diff
changeset
|
103 |
|
969
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
104 |
(* TODO: replace with proper thing *) |
966
8ba35ce16f7e
Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
962
diff
changeset
|
105 |
ML {* |
970
c16135580a45
Some processing of variables in constructors to get free variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
969
diff
changeset
|
106 |
fun is_atom ty = ty = "name" |
966
8ba35ce16f7e
Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
962
diff
changeset
|
107 |
*} |
8ba35ce16f7e
Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
962
diff
changeset
|
108 |
|
970
c16135580a45
Some processing of variables in constructors to get free variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
969
diff
changeset
|
109 |
ML {* @{term "A \<union> B"} *} |
c16135580a45
Some processing of variables in constructors to get free variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
969
diff
changeset
|
110 |
|
c16135580a45
Some processing of variables in constructors to get free variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
969
diff
changeset
|
111 |
(* TODO: After variant_frees, check that the new names correspond to the ones given by user *) |
954
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
112 |
ML {* |
970
c16135580a45
Some processing of variables in constructors to get free variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
969
diff
changeset
|
113 |
fun fv_constr lthy prefix dt_names (((name, anno_tys), bds), syn) = |
969
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
114 |
let |
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
115 |
fun prepare_name (NONE, ty) = ("", ty) |
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
116 |
| prepare_name (SOME n, ty) = (n, ty); |
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
117 |
val vars = map Free (Variable.variant_frees lthy [] (map prepare_name anno_tys)); |
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
118 |
val var_strs = map (Syntax.string_of_term lthy) vars; |
970
c16135580a45
Some processing of variables in constructors to get free variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
969
diff
changeset
|
119 |
fun fv_var (t as Free (s, (ty as Type (tyS, _)))) = |
c16135580a45
Some processing of variables in constructors to get free variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
969
diff
changeset
|
120 |
if is_atom tyS then HOLogic.mk_set ty [t] else |
c16135580a45
Some processing of variables in constructors to get free variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
969
diff
changeset
|
121 |
if tyS mem dt_names then Free ("fv_" ^ tyS, dummyT) $ t else |
c16135580a45
Some processing of variables in constructors to get free variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
969
diff
changeset
|
122 |
@{term "{}"} |
c16135580a45
Some processing of variables in constructors to get free variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
969
diff
changeset
|
123 |
val fv_eq = |
c16135580a45
Some processing of variables in constructors to get free variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
969
diff
changeset
|
124 |
if null vars then HOLogic.mk_set @{typ atom} [] |
c16135580a45
Some processing of variables in constructors to get free variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
969
diff
changeset
|
125 |
else foldr1 (HOLogic.mk_binop @{const_name union}) (map fv_var vars) |
c16135580a45
Some processing of variables in constructors to get free variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
969
diff
changeset
|
126 |
val fv_eq_str = Syntax.string_of_term lthy fv_eq |
969
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
127 |
in |
970
c16135580a45
Some processing of variables in constructors to get free variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
969
diff
changeset
|
128 |
prefix ^ " (" ^ (Binding.name_of name) ^ " " ^ (implode (separate " " var_strs)) ^ ") = " ^ fv_eq_str |
969
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
129 |
end |
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
130 |
*} |
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
131 |
|
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
132 |
ML {* |
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
133 |
fun single_dt lthy (((s2, s3), syn), constrs) = |
954
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
134 |
["constructor declaration"] |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
135 |
@ ["type arguments: "] @ s2 |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
136 |
@ ["datatype name: ", Binding.name_of s3] |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
137 |
@ ["typ mixfix: ", Pretty.string_of (Syntax.pretty_mixfix syn)] |
969
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
138 |
@ ["constructors: "] @ (map (print_constr lthy) constrs) |
954
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
139 |
|> separate "\n" |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
140 |
|> implode |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
141 |
*} |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
142 |
|
966
8ba35ce16f7e
Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
962
diff
changeset
|
143 |
ML {* |
970
c16135580a45
Some processing of variables in constructors to get free variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
969
diff
changeset
|
144 |
fun fv_dt lthy dt_names (((s2, s3), syn), constrs) = |
c16135580a45
Some processing of variables in constructors to get free variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
969
diff
changeset
|
145 |
map (fv_constr lthy ("fv_" ^ Binding.name_of s3) dt_names) constrs |
966
8ba35ce16f7e
Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
962
diff
changeset
|
146 |
|> separate "\n" |
8ba35ce16f7e
Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
962
diff
changeset
|
147 |
|> implode |
8ba35ce16f7e
Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
962
diff
changeset
|
148 |
*} |
8ba35ce16f7e
Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
962
diff
changeset
|
149 |
|
969
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
150 |
ML {* fun single_fun_eq lthy (at, s) = |
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
151 |
["function equations: ", (Syntax.string_of_term lthy s)] |
954
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
152 |
|> separate "\n" |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
153 |
|> implode |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
154 |
*} |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
155 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
156 |
ML {* fun single_fun_fix (b, _, _) = |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
157 |
["functions: ", Binding.name_of b] |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
158 |
|> separate "\n" |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
159 |
|> implode |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
160 |
*} |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
161 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
162 |
ML {* |
970
c16135580a45
Some processing of variables in constructors to get free variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
969
diff
changeset
|
163 |
fun dt_name (((s2, s3), syn), constrs) = Binding.name_of s3 |
c16135580a45
Some processing of variables in constructors to get free variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
969
diff
changeset
|
164 |
*} |
c16135580a45
Some processing of variables in constructors to get free variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
969
diff
changeset
|
165 |
|
c16135580a45
Some processing of variables in constructors to get free variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
969
diff
changeset
|
166 |
ML {* |
969
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
167 |
fun print_dts (dts, (funs, feqs)) lthy = |
954
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
168 |
let |
969
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
169 |
val _ = warning (implode (separate "\n" (map (single_dt lthy) dts))); |
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
170 |
val _ = warning (implode (separate "\n" (map single_fun_fix funs))); |
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
171 |
val _ = warning (implode (separate "\n" (map (single_fun_eq lthy) feqs))); |
970
c16135580a45
Some processing of variables in constructors to get free variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
969
diff
changeset
|
172 |
val _ = warning (implode (separate "\n" (map (fv_dt lthy (map dt_name dts)) dts))); |
954
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
173 |
in |
968
362402adfb88
Undid the parsing, as it is not possible with thy->lthy interaction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
966
diff
changeset
|
174 |
lthy |
954
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
175 |
end |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
176 |
*} |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
177 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
178 |
ML {* |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
179 |
parser; |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
180 |
Datatype.datatype_cmd; |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
181 |
*} |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
182 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
183 |
ML {* |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
184 |
fun transp_ty_arg (anno, ty) = ty |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
185 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
186 |
fun transp_constr (((constr_name, ty_args), bind), mx) = |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
187 |
(constr_name, map transp_ty_arg ty_args, mx) |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
188 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
189 |
fun transp_dt (((ty_args, ty_name), mx), constrs) = |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
190 |
(ty_args, ty_name, mx, map transp_constr constrs) |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
191 |
*} |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
192 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
193 |
ML {* |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
194 |
fun dt_defn dts thy = |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
195 |
let |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
196 |
val names = map (fn (_, s, _, _) => Binding.name_of s) dts |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
197 |
val thy' = Datatype.datatype_cmd names dts thy |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
198 |
in |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
199 |
thy' |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
200 |
end |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
201 |
*} |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
202 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
203 |
ML {* |
968
362402adfb88
Undid the parsing, as it is not possible with thy->lthy interaction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
966
diff
changeset
|
204 |
fun fn_defn [] [] lthy = lthy |
362402adfb88
Undid the parsing, as it is not possible with thy->lthy interaction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
966
diff
changeset
|
205 |
| fn_defn funs feqs lthy = |
969
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
206 |
Function_Fun.add_fun Function_Common.default_config funs feqs false lthy |
954
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
207 |
*} |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
208 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
209 |
ML {* |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
210 |
fun nominal_datatype2_cmd (dts, (funs, feqs)) thy = |
969
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
211 |
let |
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
212 |
val thy' = dt_defn (map transp_dt dts) thy |
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
213 |
val lthy = Theory_Target.init NONE thy' |
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
214 |
fun parse_fun (b, NONE, m) = (b, NONE, m) |
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
215 |
| parse_fun (b, SOME ty, m) = (b, SOME (Syntax.read_typ lthy ty), m); |
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
216 |
val funs = map parse_fun funs |
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
217 |
fun parse_feq (b, t) = (b, Syntax.read_prop lthy t); |
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
218 |
val feqs = map parse_feq feqs |
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
219 |
fun parse_anno_ty (anno, ty) = (anno, Syntax.read_typ lthy ty); |
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
220 |
fun parse_constr (((constr_name, ty_args), bind), mx) = |
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
221 |
(((constr_name, map parse_anno_ty ty_args), bind), mx); |
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
222 |
fun parse_dt (((ty_args, ty_name), mx), constrs) = |
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
223 |
(((ty_args, ty_name), mx), map parse_constr constrs); |
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
224 |
val dts = map parse_dt dts |
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
225 |
in |
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
226 |
fn_defn funs feqs lthy |
966
8ba35ce16f7e
Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
962
diff
changeset
|
227 |
|> print_dts (dts, (funs, feqs)) |
8ba35ce16f7e
Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
962
diff
changeset
|
228 |
|> Local_Theory.exit_global |
969
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
968
diff
changeset
|
229 |
end |
966
8ba35ce16f7e
Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
962
diff
changeset
|
230 |
*} |
954
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
231 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
232 |
(* Command Keyword *) |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
233 |
ML {* |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
234 |
let |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
235 |
val kind = OuterKeyword.thy_decl |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
236 |
in |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
237 |
OuterSyntax.command "nominal_datatype" "test" kind |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
238 |
(parser >> (Toplevel.theory o nominal_datatype2_cmd)) |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
239 |
end |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
240 |
*} |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
241 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
242 |
text {* example 1 *} |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
243 |
nominal_datatype lam = |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
244 |
VAR "name" |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
245 |
| APP "lam" "lam" |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
246 |
| LET bp::"bp" t::"lam" bind "bi bp" in t ("Let _ in _" [100,100] 100) |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
247 |
and bp = |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
248 |
BP "name" "lam" ("_ ::= _" [100,100] 100) |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
249 |
binder |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
250 |
bi::"bp \<Rightarrow> name set" |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
251 |
where |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
252 |
"bi (BP x t) = {x}" |
961
0f88e04eb486
Variable takes a 'name'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
954
diff
changeset
|
253 |
print_theorems |
954
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
254 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
255 |
text {* examples 2 *} |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
256 |
nominal_datatype trm = |
961
0f88e04eb486
Variable takes a 'name'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
954
diff
changeset
|
257 |
Var "name" |
954
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
258 |
| App "trm" "trm" |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
259 |
| Lam x::"name" t::"trm" bindset x in t |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
260 |
| Let p::"pat" "trm" t::"trm" bind "f p" in t |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
261 |
and pat = |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
262 |
PN |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
263 |
| PS "name" |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
264 |
| PD "name \<times> name" |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
265 |
binder |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
266 |
f::"pat \<Rightarrow> name set" |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
267 |
where |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
268 |
"f PN = {}" |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
269 |
| "f (PS x) = {x}" |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
270 |
| "f (PD (x,y)) = {x,y}" |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
271 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
272 |
nominal_datatype trm2 = |
962
8a16d6c45720
Another string in the specification.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
961
diff
changeset
|
273 |
Var2 "name" |
954
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
274 |
| App2 "trm2" "trm2" |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
275 |
| Lam2 x::"name" t::"trm2" bindset x in t |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
276 |
| Let2 p::"pat2" "trm2" t::"trm2" bind "f2 p" in t |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
277 |
and pat2 = |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
278 |
PN2 |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
279 |
| PS2 "name" |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
280 |
| PD2 "pat2 \<times> pat2" |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
281 |
binder |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
282 |
f2::"pat2 \<Rightarrow> name set" |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
283 |
where |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
284 |
"f2 PN2 = {}" |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
285 |
| "f2 (PS2 x) = {x}" |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
286 |
| "f2 (PD2 (p1, p2)) = (f2 p1) \<union> (f2 p2)" |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
287 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
288 |
text {* example type schemes *} |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
289 |
datatype ty = |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
290 |
Var "name" |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
291 |
| Fun "ty" "ty" |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
292 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
293 |
nominal_datatype tyS = |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
294 |
All xs::"name list" ty::"ty" bindset xs in ty |
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
295 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
296 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
297 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
298 |
|
961
0f88e04eb486
Variable takes a 'name'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
954
diff
changeset
|
299 |
end |