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>
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>
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>
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>
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>
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>
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 |
|
972
|
102 |
(* TODO: replace with the proper thing *)
|
966
|
103 |
ML {*
|
972
|
104 |
fun is_atom ty = ty = "Test.name"
|
966
|
105 |
*}
|
|
106 |
|
979
|
107 |
ML {*
|
|
108 |
fun fv_bds s bds set =
|
|
109 |
case bds of
|
|
110 |
[] => set
|
|
111 |
| B (s1, s2) :: tl =>
|
|
112 |
if s2 = s then
|
|
113 |
fv_bds s tl (HOLogic.mk_binop @{const_name "HOL.minus_class.minus"} (set, HOLogic.mk_set @{typ name} [Free (s1, @{typ name})]))
|
|
114 |
else fv_bds s tl set
|
|
115 |
| BS (s1, s2) :: tl =>
|
|
116 |
(* TODO: This is just a copy *)
|
|
117 |
if s2 = s then
|
|
118 |
fv_bds s tl (HOLogic.mk_binop @{const_name "HOL.minus_class.minus"} (set, HOLogic.mk_set @{typ name} [Free (s1, @{typ name})]))
|
|
119 |
else fv_bds s tl set
|
|
120 |
*}
|
|
121 |
|
972
|
122 |
(* TODO: After variant_frees, check that the new names correspond to the ones given by user
|
|
123 |
so that 'bind' is matched with variables correctly *)
|
954
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
124 |
ML {*
|
970
c16135580a45
Some processing of variables in constructors to get free variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
125 |
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>
diff
changeset
|
126 |
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>
diff
changeset
|
127 |
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>
diff
changeset
|
128 |
| 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>
diff
changeset
|
129 |
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>
diff
changeset
|
130 |
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>
diff
changeset
|
131 |
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>
diff
changeset
|
132 |
if is_atom tyS then HOLogic.mk_set ty [t] else
|
979
|
133 |
if (Long_Name.base_name tyS) mem dt_names then
|
1034
|
134 |
fv_bds s bds (Free ("rfv_" ^ (Long_Name.base_name tyS), Type ("fun", [dummyT, @{typ "name set"}])) $ t) else
|
973
|
135 |
HOLogic.mk_set @{typ name} []
|
979
|
136 |
val fv_eq =
|
973
|
137 |
if null vars then HOLogic.mk_set @{typ name} []
|
970
c16135580a45
Some processing of variables in constructors to get free variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
138 |
else foldr1 (HOLogic.mk_binop @{const_name union}) (map fv_var vars)
|
973
|
139 |
val fv_eq_str = Syntax.string_of_term lthy (Syntax.check_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>
diff
changeset
|
140 |
in
|
970
c16135580a45
Some processing of variables in constructors to get free variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
141 |
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>
diff
changeset
|
142 |
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>
diff
changeset
|
143 |
*}
|
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
144 |
|
cc5d18c78446
Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
145 |
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>
diff
changeset
|
146 |
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
|
147 |
["constructor declaration"]
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
148 |
@ ["type arguments: "] @ s2
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
149 |
@ ["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
|
150 |
@ ["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>
diff
changeset
|
151 |
@ ["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
|
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 |
|
966
|
156 |
ML {*
|
970
c16135580a45
Some processing of variables in constructors to get free variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
157 |
fun fv_dt lthy dt_names (((s2, s3), syn), constrs) =
|
1034
|
158 |
map (fv_constr lthy ("rfv_" ^ Binding.name_of s3) dt_names) constrs
|
966
|
159 |
|> separate "\n"
|
|
160 |
|> implode
|
|
161 |
*}
|
|
162 |
|
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>
diff
changeset
|
163 |
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>
diff
changeset
|
164 |
["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
|
165 |
|> separate "\n"
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
166 |
|> implode
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
167 |
*}
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
168 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
169 |
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
|
170 |
["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
|
171 |
|> separate "\n"
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
172 |
|> implode
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
173 |
*}
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
174 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
175 |
ML {*
|
970
c16135580a45
Some processing of variables in constructors to get free variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
176 |
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>
diff
changeset
|
177 |
*}
|
c16135580a45
Some processing of variables in constructors to get free variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
178 |
|
c16135580a45
Some processing of variables in constructors to get free variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
179 |
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>
diff
changeset
|
180 |
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
|
181 |
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>
diff
changeset
|
182 |
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>
diff
changeset
|
183 |
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>
diff
changeset
|
184 |
val _ = warning (implode (separate "\n" (map (single_fun_eq lthy) feqs)));
|
978
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
185 |
in
|
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
186 |
()
|
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
187 |
end
|
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
188 |
*}
|
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
189 |
|
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
190 |
ML {*
|
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
191 |
fun fv_dts (dts, (funs, feqs)) lthy =
|
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
192 |
let
|
970
c16135580a45
Some processing of variables in constructors to get free variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
193 |
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
|
194 |
in
|
968
362402adfb88
Undid the parsing, as it is not possible with thy->lthy interaction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
195 |
lthy
|
978
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
196 |
end
|
954
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
197 |
*}
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
198 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
199 |
ML {*
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
200 |
parser;
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
201 |
Datatype.datatype_cmd;
|
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 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
204 |
ML {*
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
205 |
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
|
206 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
207 |
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
|
208 |
(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
|
209 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
210 |
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
|
211 |
(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
|
212 |
*}
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
213 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
214 |
ML {*
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
215 |
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
|
216 |
let
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
217 |
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
|
218 |
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
|
219 |
in
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
220 |
thy'
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
221 |
end
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
222 |
*}
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
223 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
224 |
ML {*
|
968
362402adfb88
Undid the parsing, as it is not possible with thy->lthy interaction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
225 |
fun fn_defn [] [] lthy = lthy
|
362402adfb88
Undid the parsing, as it is not possible with thy->lthy interaction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
226 |
| fn_defn funs feqs lthy =
|
978
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
227 |
Function_Fun.add_fun Function_Common.default_config funs feqs false lthy
|
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
228 |
*}
|
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
229 |
|
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
230 |
ML {*
|
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
231 |
fun fn_defn_cmd [] [] lthy = lthy
|
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
232 |
| fn_defn_cmd funs feqs lthy =
|
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
233 |
Function_Fun.add_fun_cmd 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
|
234 |
*}
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
235 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
236 |
ML {*
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
237 |
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>
diff
changeset
|
238 |
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>
diff
changeset
|
239 |
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>
diff
changeset
|
240 |
val lthy = Theory_Target.init NONE thy'
|
978
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
241 |
val lthy' = fn_defn_cmd 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>
diff
changeset
|
242 |
fun parse_fun (b, NONE, m) = (b, NONE, m)
|
978
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
243 |
| parse_fun (b, SOME ty, m) = (b, SOME (Syntax.read_typ lthy' ty), m);
|
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>
diff
changeset
|
244 |
val funs = map parse_fun funs
|
978
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
245 |
fun parse_feq (b, t) = (b, Syntax.read_prop lthy' t);
|
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>
diff
changeset
|
246 |
val feqs = map parse_feq feqs
|
978
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
247 |
fun parse_anno_ty (anno, ty) = (anno, Syntax.read_typ lthy' ty);
|
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>
diff
changeset
|
248 |
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>
diff
changeset
|
249 |
(((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>
diff
changeset
|
250 |
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>
diff
changeset
|
251 |
(((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>
diff
changeset
|
252 |
val dts = map parse_dt dts
|
978
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
253 |
val _ = print_dts (dts, (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>
diff
changeset
|
254 |
in
|
978
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
255 |
fv_dts (dts, (funs, feqs)) lthy
|
966
|
256 |
|> 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>
diff
changeset
|
257 |
end
|
966
|
258 |
*}
|
954
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
259 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
260 |
(* Command Keyword *)
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
261 |
ML {*
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
262 |
let
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
263 |
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
|
264 |
in
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
265 |
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
|
266 |
(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
|
267 |
end
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
268 |
*}
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
269 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
270 |
text {* example 1 *}
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
271 |
nominal_datatype lam =
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
272 |
VAR "name"
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
273 |
| APP "lam" "lam"
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
274 |
| 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
|
275 |
and bp =
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
276 |
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
|
277 |
binder
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
278 |
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
|
279 |
where
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
280 |
"bi (BP x t) = {x}"
|
961
|
281 |
print_theorems
|
954
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
282 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
283 |
text {* examples 2 *}
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
284 |
nominal_datatype trm =
|
961
|
285 |
Var "name"
|
954
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
286 |
| App "trm" "trm"
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
287 |
| 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
|
288 |
| 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
|
289 |
and pat =
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
290 |
PN
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
291 |
| PS "name"
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
292 |
| 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
|
293 |
binder
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
294 |
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
|
295 |
where
|
978
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
296 |
"f PN = {}"
|
954
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
297 |
| "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
|
298 |
| "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
|
299 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
300 |
nominal_datatype trm2 =
|
962
|
301 |
Var2 "name"
|
954
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
302 |
| App2 "trm2" "trm2"
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
303 |
| 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
|
304 |
| 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
|
305 |
and pat2 =
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
306 |
PN2
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
307 |
| PS2 "name"
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
308 |
| 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
|
309 |
binder
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
310 |
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
|
311 |
where
|
978
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
312 |
"f2 PN2 = {}"
|
954
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
313 |
| "f2 (PS2 x) = {x}"
|
978
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
314 |
| "f2 (PD2 (p1, p2)) = (f2 p1) \<union> (f2 p2)"
|
954
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
315 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
316 |
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
|
317 |
datatype ty =
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
318 |
Var "name"
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
319 |
| Fun "ty" "ty"
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
320 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
321 |
nominal_datatype tyS =
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
322 |
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
|
323 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
324 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
325 |
|
c009d2535896
very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
326 |
|
961
|
327 |
end
|