| author | Christian Urban <urbanc@in.tum.de> | 
| Mon, 01 Mar 2010 19:23:08 +0100 | |
| changeset 1290 | a7e7ffb7f362 | 
| parent 1289 | 02eb0f600630 | 
| child 1293 | dca51a1f0c0d | 
| permissions | -rw-r--r-- | 
| 1261 
853abc14c5c6
added IsaMakefile...but so far included only a test for the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
changeset | 1 | theory Parser | 
| 954 
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 | |
| 1194 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 7 | section{* Interface for nominal_datatype *}
 | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 8 | |
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 9 | text {*
 | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 10 | |
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 11 | Nominal-Datatype-part: | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 12 | |
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 13 | |
| 1266 | 14 | 1nd Arg: (string list * binding * mixfix * (binding * typ list * mixfix) list) list | 
| 1194 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 15 | ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^ ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^ | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 16 | type(s) to be defined constructors list | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 17 | (ty args, name, syn) (name, typs, syn) | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 18 | |
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 19 | Binder-Function-part: | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 20 | |
| 1266 | 21 | 2rd Arg: (binding * typ option * mixfix) list | 
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 22 | ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^ | 
| 1194 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 23 | binding function(s) | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 24 | to be defined | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 25 | (name, type, syn) | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 26 | |
| 1266 | 27 | 3th Arg: term list | 
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 28 | ^^^^^^^^^ | 
| 1194 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 29 | the equations of the binding functions | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 30 | (Trueprop equations) | 
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 31 | *} | 
| 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 32 | |
| 1283 
6a133abb7633
generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
 Christian Urban <urbanc@in.tum.de> parents: 
1266diff
changeset | 33 | ML {*
 | 
| 
6a133abb7633
generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
 Christian Urban <urbanc@in.tum.de> parents: 
1266diff
changeset | 34 | |
| 
6a133abb7633
generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
 Christian Urban <urbanc@in.tum.de> parents: 
1266diff
changeset | 35 | *} | 
| 
6a133abb7633
generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
 Christian Urban <urbanc@in.tum.de> parents: 
1266diff
changeset | 36 | |
| 1194 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 37 | text {*****************************************************}
 | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 38 | ML {* 
 | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 39 | (* nominal datatype parser *) | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 40 | local | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 41 | structure P = OuterParse | 
| 1266 | 42 | |
| 43 | fun tuple ((x, y, z), u) = (x, y, z, u) | |
| 1283 
6a133abb7633
generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
 Christian Urban <urbanc@in.tum.de> parents: 
1266diff
changeset | 44 | fun tswap (((x, y), z), u) = (x, y, u, z) | 
| 1194 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 45 | in | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 46 | |
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 47 | val _ = OuterKeyword.keyword "bind" | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 48 | val anno_typ = Scan.option (P.name --| P.$$$ "::") -- P.typ | 
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 49 | |
| 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 50 | (* binding specification *) | 
| 1266 | 51 | (* maybe use and_list *) | 
| 954 
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 = | 
| 1283 
6a133abb7633
generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
 Christian Urban <urbanc@in.tum.de> parents: 
1266diff
changeset | 53 | P.enum "," ((P.$$$ "bind" |-- P.term) -- (P.$$$ "in" |-- P.name) >> swap) | 
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 54 | |
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 55 | val constr_parser = | 
| 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 56 | P.binding -- Scan.repeat anno_typ | 
| 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 57 | |
| 954 
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 | val dt_parser = | 
| 1283 
6a133abb7633
generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
 Christian Urban <urbanc@in.tum.de> parents: 
1266diff
changeset | 60 | (P.type_args -- P.binding -- P.opt_mixfix >> P.triple1) -- | 
| 
6a133abb7633
generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
 Christian Urban <urbanc@in.tum.de> parents: 
1266diff
changeset | 61 | (P.$$$ "=" |-- P.enum1 "|" (constr_parser -- bind_parser -- P.opt_mixfix >> tswap)) >> tuple | 
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 62 | |
| 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 63 | (* function equation parser *) | 
| 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 64 | val fun_parser = | 
| 1194 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 65 | Scan.optional (P.$$$ "binder" |-- P.fixes -- SpecParse.where_alt_specs) ([],[]) | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 66 | |
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 67 | (* main parser *) | 
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 68 | val main_parser = | 
| 1266 | 69 | (P.and_list1 dt_parser) -- fun_parser >> P.triple2 | 
| 1194 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 70 | |
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 71 | end | 
| 954 
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 | |
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 74 | (* adds "_raw" to the end of constants and types *) | 
| 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 75 | ML {*
 | 
| 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 76 | fun add_raw s = s ^ "_raw" | 
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 77 | fun add_raws ss = map add_raw ss | 
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 78 | fun raw_bind bn = Binding.suffix_name "_raw" bn | 
| 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 79 | |
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 80 | fun replace_str ss s = | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 81 | case (AList.lookup (op=) ss s) of | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 82 | SOME s' => s' | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 83 | | NONE => s | 
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 84 | |
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 85 | fun replace_typ ty_ss (Type (a, Ts)) = Type (replace_str ty_ss a, map (replace_typ ty_ss) Ts) | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 86 | | replace_typ ty_ss T = T | 
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 87 | |
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 88 | fun raw_dts ty_ss dts = | 
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 89 | let | 
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 90 | |
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 91 | fun raw_dts_aux1 (bind, tys, mx) = | 
| 1289 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 92 | (raw_bind bind, map (replace_typ ty_ss) tys, mx) | 
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 93 | |
| 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 94 | fun raw_dts_aux2 (ty_args, bind, mx, constrs) = | 
| 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 95 | (ty_args, raw_bind bind, mx, map raw_dts_aux1 constrs) | 
| 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 96 | in | 
| 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 97 | map raw_dts_aux2 dts | 
| 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 98 | end | 
| 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 99 | |
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 100 | fun replace_aterm trm_ss (Const (a, T)) = Const (replace_str trm_ss a, T) | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 101 | | replace_aterm trm_ss (Free (a, T)) = Free (replace_str trm_ss a, T) | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 102 | | replace_aterm trm_ss trm = trm | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 103 | |
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 104 | fun replace_term trm_ss ty_ss trm = | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 105 | trm |> Term.map_aterms (replace_aterm trm_ss) |> map_types (replace_typ ty_ss) | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 106 | *} | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 107 | |
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 108 | ML {*
 | 
| 1289 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 109 | fun get_cnstrs dts = | 
| 1290 
a7e7ffb7f362
modified for new binding format - hope it is the intended one
 Christian Urban <urbanc@in.tum.de> parents: 
1289diff
changeset | 110 | map (fn (_, _, _, constrs) => constrs) dts | 
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 111 | |
| 1289 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 112 | fun get_typed_cnstrs dts = | 
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 113 | flat (map (fn (_, bn, _, constrs) => | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 114 | (map (fn (bn', _, _) => (Binding.name_of bn, Binding.name_of bn')) constrs)) dts) | 
| 1232 
35bc8f7b66f8
parsing of function definitions almost works now; still an error with undefined constants
 Christian Urban <urbanc@in.tum.de> parents: 
1228diff
changeset | 115 | |
| 1289 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 116 | fun get_cnstr_strs dts = | 
| 1290 
a7e7ffb7f362
modified for new binding format - hope it is the intended one
 Christian Urban <urbanc@in.tum.de> parents: 
1289diff
changeset | 117 | map (fn (bn, _, _) => Binding.name_of bn) (flat (get_cnstrs dts)) | 
| 1232 
35bc8f7b66f8
parsing of function definitions almost works now; still an error with undefined constants
 Christian Urban <urbanc@in.tum.de> parents: 
1228diff
changeset | 118 | |
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 119 | fun get_bn_fun_strs bn_funs = | 
| 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 120 | map (fn (bn_fun, _, _) => Binding.name_of bn_fun) bn_funs | 
| 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 121 | *} | 
| 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 122 | |
| 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 123 | ML {*
 | 
| 1289 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 124 | fun rawify_dts dt_names dts dts_env = | 
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 125 | let | 
| 1289 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 126 | val raw_dts = raw_dts dts_env dts | 
| 1287 
8557af71724e
slight simplification of the raw-decl generation
 Christian Urban <urbanc@in.tum.de> parents: 
1285diff
changeset | 127 | val raw_dt_names = add_raws dt_names | 
| 1228 
c179ad9d2446
declarartion of the raw datatype already works; raw binding functions throw an exception about mutual recursive types
 Christian Urban <urbanc@in.tum.de> parents: 
1223diff
changeset | 128 | in | 
| 1287 
8557af71724e
slight simplification of the raw-decl generation
 Christian Urban <urbanc@in.tum.de> parents: 
1285diff
changeset | 129 | (raw_dt_names, raw_dts) | 
| 1228 
c179ad9d2446
declarartion of the raw datatype already works; raw binding functions throw an exception about mutual recursive types
 Christian Urban <urbanc@in.tum.de> parents: 
1223diff
changeset | 130 | end | 
| 
c179ad9d2446
declarartion of the raw datatype already works; raw binding functions throw an exception about mutual recursive types
 Christian Urban <urbanc@in.tum.de> parents: 
1223diff
changeset | 131 | *} | 
| 
c179ad9d2446
declarartion of the raw datatype already works; raw binding functions throw an exception about mutual recursive types
 Christian Urban <urbanc@in.tum.de> parents: 
1223diff
changeset | 132 | |
| 
c179ad9d2446
declarartion of the raw datatype already works; raw binding functions throw an exception about mutual recursive types
 Christian Urban <urbanc@in.tum.de> parents: 
1223diff
changeset | 133 | ML {*
 | 
| 1289 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 134 | fun rawify_bn_funs dts_env cnstrs_env bn_fun_env bn_funs bn_eqs = | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 135 | let | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 136 | val bn_funs' = map (fn (bn, ty, mx) => | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 137 | (raw_bind bn, replace_typ dts_env ty, mx)) bn_funs | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 138 | |
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 139 | val bn_eqs' = map (fn (attr, trm) => | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 140 | (attr, replace_term (cnstrs_env @ bn_fun_env) dts_env trm)) bn_eqs | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 141 | in | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 142 | (bn_funs', bn_eqs') | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 143 | end | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 144 | *} | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 145 | |
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 146 | ML {* 
 | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 147 | fun add_primrec_wrapper funs eqs lthy = | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 148 | if null funs then (([], []), lthy) | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 149 | else | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 150 | let | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 151 | val eqs' = map (fn (_, eq) => (Attrib.empty_binding, eq)) eqs | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 152 | val funs' = map (fn (bn, ty, mx) => (bn, SOME ty, mx)) funs | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 153 | in | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 154 | Primrec.add_primrec funs' eqs' lthy | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 155 | end | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 156 | *} | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 157 | |
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 158 | |
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 159 | ML {* 
 | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 160 | fun nominal_datatype2 dts bn_funs bn_eqs binds lthy = | 
| 1228 
c179ad9d2446
declarartion of the raw datatype already works; raw binding functions throw an exception about mutual recursive types
 Christian Urban <urbanc@in.tum.de> parents: 
1223diff
changeset | 161 | let | 
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 162 | val thy = ProofContext.theory_of lthy | 
| 1289 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 163 | val thy_name = Context.theory_name thy | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 164 | |
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 165 | val conf = Datatype.default_config | 
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 166 | |
| 1289 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 167 | val dt_names = map (fn (_, s, _, _) => Binding.name_of s) dts | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 168 | val dt_full_names = map (Long_Name.qualify thy_name) dt_names | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 169 | val dt_full_names' = add_raws dt_full_names | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 170 | val dts_env = dt_full_names ~~ dt_full_names' | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 171 | |
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 172 | val cnstrs = get_cnstr_strs dts | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 173 | val cnstrs_ty = get_typed_cnstrs dts | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 174 | val cnstrs_full_names = map (Long_Name.qualify thy_name) cnstrs | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 175 | val cnstrs_full_names' = map (fn (x, y) => Long_Name.qualify thy_name | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 176 | (Long_Name.qualify (add_raw x) (add_raw y))) cnstrs_ty | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 177 | val cnstrs_env = cnstrs_full_names ~~ cnstrs_full_names' | 
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 178 | |
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 179 | val bn_fun_strs = get_bn_fun_strs bn_funs | 
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 180 | val bn_fun_strs' = add_raws bn_fun_strs | 
| 1266 | 181 | val bn_fun_env = bn_fun_strs ~~ bn_fun_strs' | 
| 1289 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 182 | |
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 183 | val (raw_dt_names, raw_dts) = rawify_dts dt_names dts dts_env | 
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 184 | |
| 1289 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 185 | val (raw_bn_funs, raw_bn_eqs) = rawify_bn_funs dts_env cnstrs_env bn_fun_env bn_funs bn_eqs | 
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 186 | in | 
| 1289 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 187 | lthy | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 188 | |> Local_Theory.theory_result (Datatype.add_datatype conf raw_dt_names raw_dts) | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 189 | ||>> add_primrec_wrapper raw_bn_funs raw_bn_eqs | 
| 978 
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
973diff
changeset | 190 | end | 
| 
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
973diff
changeset | 191 | *} | 
| 
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
973diff
changeset | 192 | |
| 1284 | 193 | ML {* 
 | 
| 194 | (* parsing the datatypes and declaring *) | |
| 195 | (* constructors in the local theory *) | |
| 196 | fun prepare_dts dt_strs lthy = | |
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 197 | let | 
| 1283 
6a133abb7633
generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
 Christian Urban <urbanc@in.tum.de> parents: 
1266diff
changeset | 198 | val thy = ProofContext.theory_of lthy | 
| 1284 | 199 | |
| 200 | fun mk_type full_tname tvrs = | |
| 201 | Type (full_tname, map (fn a => TVar ((a, 0), [])) tvrs) | |
| 1283 
6a133abb7633
generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
 Christian Urban <urbanc@in.tum.de> parents: 
1266diff
changeset | 202 | |
| 1284 | 203 | fun prep_cnstr lthy full_tname tvs (cname, anno_tys, mx, _) = | 
| 204 | let | |
| 205 | val tys = map (Syntax.read_typ lthy o snd) anno_tys | |
| 206 | val ty = mk_type full_tname tvs | |
| 207 | in | |
| 208 | ((cname, tys ---> ty, mx), (cname, tys, mx)) | |
| 209 | end | |
| 210 | ||
| 211 | fun prep_dt lthy (tvs, tname, mx, cnstrs) = | |
| 212 | let | |
| 213 | val full_tname = Sign.full_name thy tname | |
| 214 | val (cnstrs', cnstrs'') = | |
| 215 | split_list (map (prep_cnstr lthy full_tname tvs) cnstrs) | |
| 216 | in | |
| 217 | (cnstrs', (tvs, tname, mx, cnstrs'')) | |
| 218 | end | |
| 1266 | 219 | |
| 1284 | 220 | val (cnstrs, dts) = | 
| 221 | split_list (map (prep_dt lthy) dt_strs) | |
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 222 | in | 
| 1284 | 223 | lthy | 
| 224 | |> Local_Theory.theory (Sign.add_consts_i (flat cnstrs)) | |
| 225 | |> pair dts | |
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 226 | end | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 227 | *} | 
| 1232 
35bc8f7b66f8
parsing of function definitions almost works now; still an error with undefined constants
 Christian Urban <urbanc@in.tum.de> parents: 
1228diff
changeset | 228 | |
| 1284 | 229 | ML {*
 | 
| 1285 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 230 | (* parsing the binding function specification and *) | 
| 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 231 | (* declaring the functions in the local theory *) | 
| 1284 | 232 | fun prepare_bn_funs bn_fun_strs bn_eq_strs lthy = | 
| 233 | let | |
| 234 | fun prep_bn_fun ((bn, T), mx) = (bn, T, mx) | |
| 235 | ||
| 236 | val ((bn_funs, bn_eqs), _) = | |
| 237 | Specification.read_spec bn_fun_strs bn_eq_strs lthy | |
| 238 | ||
| 239 | val bn_funs' = map prep_bn_fun bn_funs | |
| 240 | in | |
| 241 | lthy | |
| 242 | |> Local_Theory.theory (Sign.add_consts_i bn_funs') | |
| 243 | |> pair (bn_funs', bn_eqs) | |
| 244 | end | |
| 1283 
6a133abb7633
generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
 Christian Urban <urbanc@in.tum.de> parents: 
1266diff
changeset | 245 | *} | 
| 
6a133abb7633
generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
 Christian Urban <urbanc@in.tum.de> parents: 
1266diff
changeset | 246 | |
| 
6a133abb7633
generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
 Christian Urban <urbanc@in.tum.de> parents: 
1266diff
changeset | 247 | ML {*
 | 
| 1290 
a7e7ffb7f362
modified for new binding format - hope it is the intended one
 Christian Urban <urbanc@in.tum.de> parents: 
1289diff
changeset | 248 | fun find_all eq xs (k',i) = | 
| 
a7e7ffb7f362
modified for new binding format - hope it is the intended one
 Christian Urban <urbanc@in.tum.de> parents: 
1289diff
changeset | 249 | maps (fn (k, (v1, v2)) => if eq (k, k') then [(v1, v2, i)] else []) xs | 
| 1285 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 250 | *} | 
| 1284 | 251 | |
| 1285 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 252 | ML {*
 | 
| 1290 
a7e7ffb7f362
modified for new binding format - hope it is the intended one
 Christian Urban <urbanc@in.tum.de> parents: 
1289diff
changeset | 253 | (* associates every SOME with the index in the list; drops NONEs *) | 
| 1283 
6a133abb7633
generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
 Christian Urban <urbanc@in.tum.de> parents: 
1266diff
changeset | 254 | fun mk_env xs = | 
| 
6a133abb7633
generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
 Christian Urban <urbanc@in.tum.de> parents: 
1266diff
changeset | 255 | let | 
| 
6a133abb7633
generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
 Christian Urban <urbanc@in.tum.de> parents: 
1266diff
changeset | 256 | fun mapp (_: int) [] = [] | 
| 1285 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 257 | | mapp i (a :: xs) = | 
| 1283 
6a133abb7633
generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
 Christian Urban <urbanc@in.tum.de> parents: 
1266diff
changeset | 258 | case a of | 
| 
6a133abb7633
generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
 Christian Urban <urbanc@in.tum.de> parents: 
1266diff
changeset | 259 | NONE => mapp (i + 1) xs | 
| 
6a133abb7633
generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
 Christian Urban <urbanc@in.tum.de> parents: 
1266diff
changeset | 260 | | SOME x => (x, i) :: mapp (i + 1) xs | 
| 
6a133abb7633
generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
 Christian Urban <urbanc@in.tum.de> parents: 
1266diff
changeset | 261 | in mapp 0 xs end | 
| 1285 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 262 | *} | 
| 1283 
6a133abb7633
generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
 Christian Urban <urbanc@in.tum.de> parents: 
1266diff
changeset | 263 | |
| 1285 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 264 | ML {*
 | 
| 1283 
6a133abb7633
generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
 Christian Urban <urbanc@in.tum.de> parents: 
1266diff
changeset | 265 | fun env_lookup xs x = | 
| 
6a133abb7633
generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
 Christian Urban <urbanc@in.tum.de> parents: 
1266diff
changeset | 266 | case AList.lookup (op =) xs x of | 
| 
6a133abb7633
generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
 Christian Urban <urbanc@in.tum.de> parents: 
1266diff
changeset | 267 | SOME x => x | 
| 1287 
8557af71724e
slight simplification of the raw-decl generation
 Christian Urban <urbanc@in.tum.de> parents: 
1285diff
changeset | 268 |   | NONE => error ("cannot find " ^ x ^ " in the binding specification.");
 | 
| 1283 
6a133abb7633
generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
 Christian Urban <urbanc@in.tum.de> parents: 
1266diff
changeset | 269 | *} | 
| 
6a133abb7633
generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
 Christian Urban <urbanc@in.tum.de> parents: 
1266diff
changeset | 270 | |
| 978 
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
973diff
changeset | 271 | ML {*
 | 
| 1284 | 272 | fun prepare_binds dt_strs lthy = | 
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 273 | let | 
| 1287 
8557af71724e
slight simplification of the raw-decl generation
 Christian Urban <urbanc@in.tum.de> parents: 
1285diff
changeset | 274 | fun extract_annos_binds dt_strs = | 
| 1290 
a7e7ffb7f362
modified for new binding format - hope it is the intended one
 Christian Urban <urbanc@in.tum.de> parents: 
1289diff
changeset | 275 | map (map (fn (_, antys, _, bns) => (map fst antys, bns))) dt_strs | 
| 1285 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 276 | |
| 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 277 | fun prep_bn env bn_str = | 
| 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 278 | case (Syntax.read_term lthy bn_str) of | 
| 1290 
a7e7ffb7f362
modified for new binding format - hope it is the intended one
 Christian Urban <urbanc@in.tum.de> parents: 
1289diff
changeset | 279 | Free (x, _) => (NONE, env_lookup env x) | 
| 
a7e7ffb7f362
modified for new binding format - hope it is the intended one
 Christian Urban <urbanc@in.tum.de> parents: 
1289diff
changeset | 280 | | Const (a, T) $ Free (x, _) => (SOME (Const (a, T)), env_lookup env x) | 
| 1285 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 281 | | _ => error (bn_str ^ " not allowed as binding specification."); | 
| 1283 
6a133abb7633
generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
 Christian Urban <urbanc@in.tum.de> parents: 
1266diff
changeset | 282 | |
| 1290 
a7e7ffb7f362
modified for new binding format - hope it is the intended one
 Christian Urban <urbanc@in.tum.de> parents: 
1289diff
changeset | 283 | fun prep_typ env (i, opt_name) = | 
| 1285 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 284 | case opt_name of | 
| 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 285 | NONE => [] | 
| 1290 
a7e7ffb7f362
modified for new binding format - hope it is the intended one
 Christian Urban <urbanc@in.tum.de> parents: 
1289diff
changeset | 286 | | SOME x => find_all (op=) env (x,i); | 
| 1283 
6a133abb7633
generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
 Christian Urban <urbanc@in.tum.de> parents: 
1266diff
changeset | 287 | |
| 1287 
8557af71724e
slight simplification of the raw-decl generation
 Christian Urban <urbanc@in.tum.de> parents: 
1285diff
changeset | 288 | (* annos - list of annotation for each type (either NONE or SOME fo a type *) | 
| 
8557af71724e
slight simplification of the raw-decl generation
 Christian Urban <urbanc@in.tum.de> parents: 
1285diff
changeset | 289 | |
| 
8557af71724e
slight simplification of the raw-decl generation
 Christian Urban <urbanc@in.tum.de> parents: 
1285diff
changeset | 290 | fun prep_binds (annos, bind_strs) = | 
| 1284 | 291 | let | 
| 1290 
a7e7ffb7f362
modified for new binding format - hope it is the intended one
 Christian Urban <urbanc@in.tum.de> parents: 
1289diff
changeset | 292 | val env = mk_env annos (* for every label the index *) | 
| 1284 | 293 | val binds = map (fn (x, y) => (x, prep_bn env y)) bind_strs | 
| 294 | in | |
| 1290 
a7e7ffb7f362
modified for new binding format - hope it is the intended one
 Christian Urban <urbanc@in.tum.de> parents: 
1289diff
changeset | 295 | map_index (prep_typ binds) annos | 
| 1284 | 296 | end | 
| 1287 
8557af71724e
slight simplification of the raw-decl generation
 Christian Urban <urbanc@in.tum.de> parents: 
1285diff
changeset | 297 | |
| 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: 
968diff
changeset | 298 | in | 
| 1290 
a7e7ffb7f362
modified for new binding format - hope it is the intended one
 Christian Urban <urbanc@in.tum.de> parents: 
1289diff
changeset | 299 | map (map prep_binds) (extract_annos_binds (get_cnstrs dt_strs)) | 
| 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: 
968diff
changeset | 300 | end | 
| 966 
8ba35ce16f7e
Some cleaning of thy vs lthy vs context.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
962diff
changeset | 301 | *} | 
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 302 | |
| 1284 | 303 | ML {*
 | 
| 304 | fun nominal_datatype2_cmd (dt_strs, bn_fun_strs, bn_eq_strs) lthy = | |
| 305 | let | |
| 306 | fun prep_typ (tvs, tname, mx, _) = (tname, length tvs, mx) | |
| 307 | ||
| 308 | val ((dts, (bn_fun, bn_eq)), binds) = | |
| 309 | lthy | |
| 310 | |> Local_Theory.theory (Sign.add_types (map prep_typ dt_strs)) | |
| 311 | |> prepare_dts dt_strs | |
| 312 | ||>> prepare_bn_funs bn_fun_strs bn_eq_strs | |
| 313 | ||> prepare_binds dt_strs | |
| 314 | ||
| 1290 
a7e7ffb7f362
modified for new binding format - hope it is the intended one
 Christian Urban <urbanc@in.tum.de> parents: 
1289diff
changeset | 315 | val _ = tracing (PolyML.makestring binds) | 
| 1284 | 316 | in | 
| 317 | nominal_datatype2 dts bn_fun bn_eq binds lthy |> snd | |
| 318 | end | |
| 319 | *} | |
| 320 | ||
| 321 | ||
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 322 | (* Command Keyword *) | 
| 1283 
6a133abb7633
generated the "binding list" from the input; at the moment it is only printed out as tracing; does not yet include the "bind itself binders"
 Christian Urban <urbanc@in.tum.de> parents: 
1266diff
changeset | 323 | |
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 324 | ML {*
 | 
| 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 325 | let | 
| 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 326 | 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 | 327 | in | 
| 1228 
c179ad9d2446
declarartion of the raw datatype already works; raw binding functions throw an exception about mutual recursive types
 Christian Urban <urbanc@in.tum.de> parents: 
1223diff
changeset | 328 | OuterSyntax.local_theory "nominal_datatype" "test" kind | 
| 
c179ad9d2446
declarartion of the raw datatype already works; raw binding functions throw an exception about mutual recursive types
 Christian Urban <urbanc@in.tum.de> parents: 
1223diff
changeset | 329 | (main_parser >> nominal_datatype2_cmd) | 
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 330 | end | 
| 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 331 | *} | 
| 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 332 | |
| 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 333 | |
| 961 
0f88e04eb486
Variable takes a 'name'.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
954diff
changeset | 334 | end | 
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 335 | |
| 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 336 | |
| 1228 
c179ad9d2446
declarartion of the raw datatype already works; raw binding functions throw an exception about mutual recursive types
 Christian Urban <urbanc@in.tum.de> parents: 
1223diff
changeset | 337 |