| author | Christian Urban <urbanc@in.tum.de> | 
| Mon, 01 Mar 2010 07:46:50 +0100 | |
| changeset 1285 | e3ac56d3b7af | 
| parent 1284 | 212f3ab40cc2 | 
| child 1287 | 8557af71724e | 
| 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 | |
| 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 7 | |
| 1194 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 8 | section{* Interface for nominal_datatype *}
 | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 9 | |
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 10 | text {*
 | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 11 | |
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 12 | Nominal-Datatype-part: | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 13 | |
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 14 | |
| 1266 | 15 | 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 | 16 | ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^ ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^ | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 17 | type(s) to be defined constructors list | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 18 | (ty args, name, syn) (name, typs, syn) | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 19 | |
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 20 | Binder-Function-part: | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 21 | |
| 1266 | 22 | 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 | 23 | ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^ | 
| 1194 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 24 | binding function(s) | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 25 | to be defined | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 26 | (name, type, syn) | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 27 | |
| 1266 | 28 | 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 | 29 | ^^^^^^^^^ | 
| 1194 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 30 | the equations of the binding functions | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 31 | (Trueprop equations) | 
| 954 
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 | |
| 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 | 34 | 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 | 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 | *} | 
| 
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 | 37 | |
| 1194 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 38 | text {*****************************************************}
 | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 39 | ML {* 
 | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 40 | (* nominal datatype parser *) | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 41 | local | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 42 | structure P = OuterParse | 
| 1266 | 43 | |
| 44 | 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 | 45 | 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 | 46 | in | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 47 | |
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 48 | val _ = OuterKeyword.keyword "bind" | 
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 49 | 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 | 50 | |
| 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 51 | (* binding specification *) | 
| 1266 | 52 | (* 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 | 53 | 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 | 54 | 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 | 55 | |
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 56 | val constr_parser = | 
| 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 57 | 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 | 58 | |
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 59 | (* datatype parser *) | 
| 
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 = | 
| 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 | 61 | (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 | 62 | (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 | 63 | |
| 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 64 | (* function equation parser *) | 
| 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 65 | val fun_parser = | 
| 1194 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 66 | 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 | 67 | |
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 68 | (* main parser *) | 
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 69 | val main_parser = | 
| 1266 | 70 | (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 | 71 | |
| 
3d54fcc5f41a
start work with the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1087diff
changeset | 72 | end | 
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 73 | *} | 
| 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 74 | |
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 75 | (* 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 | 76 | ML {*
 | 
| 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 77 | 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 | 78 | 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 | 79 | 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 | 80 | |
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 81 | 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 | 82 | 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 | 83 | 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 | 84 | | NONE => s | 
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 85 | |
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 86 | 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 | 87 | | 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 | 88 | |
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 89 | 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 | 90 | 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 | 91 | val ty_ss' = ty_ss ~~ (add_raws 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 | 92 | |
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 93 | fun raw_dts_aux1 (bind, tys, mx) = | 
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 94 | (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 | 95 | |
| 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 96 | 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 | 97 | (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 | 98 | in | 
| 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 99 | 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 | 100 | end | 
| 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 101 | |
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 102 | 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 | 103 | | 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 | 104 | | 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 | 105 | |
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 106 | 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 | 107 | 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 | 108 | *} | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 109 | |
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 110 | ML {*
 | 
| 1232 
35bc8f7b66f8
parsing of function definitions almost works now; still an error with undefined constants
 Christian Urban <urbanc@in.tum.de> parents: 
1228diff
changeset | 111 | fun get_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 | 112 | flat (map (fn (_, _, _, constrs) => constrs) dts) | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 113 | |
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 114 | fun get_typed_constrs dts = | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 115 | 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 | 116 | (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 | 117 | |
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 118 | fun get_constr_strs 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 | 119 | map (fn (bn, _, _) => Binding.name_of bn) (get_constrs dts) | 
| 
35bc8f7b66f8
parsing of function definitions almost works now; still an error with undefined constants
 Christian Urban <urbanc@in.tum.de> parents: 
1228diff
changeset | 120 | |
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 121 | 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 | 122 | 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 | 123 | *} | 
| 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 124 | |
| 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 125 | ML {*
 | 
| 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 | 126 | fun raw_dts_decl dt_names dts lthy = | 
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 127 | 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 | 128 | val thy = ProofContext.theory_of lthy | 
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 129 | val conf = Datatype.default_config | 
| 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 | |
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 131 | val dt_names' = add_raws dt_names | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 132 | val dt_full_names = map (Sign.full_bname thy) dt_names | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 133 | val dts' = raw_dts dt_full_names 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 | 134 | in | 
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 135 | lthy | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 136 | |> Local_Theory.theory_result (Datatype.add_datatype conf dt_names' 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 | 137 | 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 | 138 | *} | 
| 
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 | 139 | |
| 
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 | 140 | ML {*
 | 
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 141 | fun raw_bn_fun_decl dt_names dts bn_funs bn_eqs 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 | 142 | 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 | 143 | val thy = ProofContext.theory_of lthy | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 144 | |
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 145 | val dt_names' = add_raws dt_names | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 146 | val dt_full_names = map (Sign.full_bname thy) dt_names | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 147 | val dt_full_names' = map (Sign.full_bname thy) dt_names' | 
| 1266 | 148 | val dt_env = dt_full_names ~~ dt_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 | 149 | |
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 150 | val ctrs_names = map (Sign.full_bname thy) (get_constr_strs dts) | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 151 | val ctrs_names' = map (fn (x, y) => (Sign.full_bname_path thy (add_raw x) (add_raw y))) | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 152 | (get_typed_constrs dts) | 
| 1266 | 153 | val ctrs_env = ctrs_names ~~ ctrs_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 | 154 | |
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 155 | 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 | 156 | val bn_fun_strs' = add_raws bn_fun_strs | 
| 1266 | 157 | val bn_fun_env = bn_fun_strs ~~ bn_fun_strs' | 
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 158 | |
| 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 | 159 | val bn_funs' = map (fn (bn, ty, mx) => | 
| 
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 | 160 | (raw_bind bn, SOME (replace_typ dt_env ty), mx)) 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 | 161 | |
| 1284 | 162 | val bn_eqs' = map (fn (_, trm) => | 
| 1266 | 163 | (Attrib.empty_binding, replace_term (ctrs_env @ bn_fun_env) dt_env trm)) bn_eqs | 
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 164 | in | 
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 165 | if null bn_eqs | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 166 | then (([], []), lthy) | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 167 | else Primrec.add_primrec bn_funs' bn_eqs' lthy | 
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 168 | end | 
| 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 169 | *} | 
| 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 170 | |
| 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 | 171 | ML {* 
 | 
| 1284 | 172 | fun nominal_datatype2 dts bn_funs bn_eqs binds lthy = | 
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 173 | let | 
| 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 | 174 | val dts_names = map (fn (_, s, _, _) => Binding.name_of s) dts | 
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 175 | 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 | 176 | 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 | 177 | |> raw_dts_decl dts_names dts | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 178 | ||>> raw_bn_fun_decl dts_names dts bn_funs bn_eqs | 
| 978 
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
973diff
changeset | 179 | end | 
| 
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
973diff
changeset | 180 | *} | 
| 
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
973diff
changeset | 181 | |
| 1284 | 182 | |
| 183 | ML {* 
 | |
| 184 | (* parsing the datatypes and declaring *) | |
| 185 | (* constructors in the local theory *) | |
| 186 | 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 | 187 | 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 | 188 | val thy = ProofContext.theory_of lthy | 
| 1284 | 189 | |
| 190 | fun mk_type full_tname tvrs = | |
| 191 | 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 | 192 | |
| 1284 | 193 | fun prep_cnstr lthy full_tname tvs (cname, anno_tys, mx, _) = | 
| 194 | let | |
| 195 | val tys = map (Syntax.read_typ lthy o snd) anno_tys | |
| 196 | val ty = mk_type full_tname tvs | |
| 197 | in | |
| 198 | ((cname, tys ---> ty, mx), (cname, tys, mx)) | |
| 199 | end | |
| 200 | ||
| 201 | fun prep_dt lthy (tvs, tname, mx, cnstrs) = | |
| 202 | let | |
| 203 | val full_tname = Sign.full_name thy tname | |
| 204 | val (cnstrs', cnstrs'') = | |
| 205 | split_list (map (prep_cnstr lthy full_tname tvs) cnstrs) | |
| 206 | in | |
| 207 | (cnstrs', (tvs, tname, mx, cnstrs'')) | |
| 208 | end | |
| 1266 | 209 | |
| 1284 | 210 | val (cnstrs, dts) = | 
| 211 | 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 | 212 | in | 
| 1284 | 213 | lthy | 
| 214 | |> Local_Theory.theory (Sign.add_consts_i (flat cnstrs)) | |
| 215 | |> 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 | 216 | end | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 217 | *} | 
| 1232 
35bc8f7b66f8
parsing of function definitions almost works now; still an error with undefined constants
 Christian Urban <urbanc@in.tum.de> parents: 
1228diff
changeset | 218 | |
| 1284 | 219 | ML {*
 | 
| 1285 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 220 | (* parsing the binding function specification and *) | 
| 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 221 | (* declaring the functions in the local theory *) | 
| 1284 | 222 | fun prepare_bn_funs bn_fun_strs bn_eq_strs lthy = | 
| 223 | let | |
| 224 | fun prep_bn_fun ((bn, T), mx) = (bn, T, mx) | |
| 225 | ||
| 226 | val ((bn_funs, bn_eqs), _) = | |
| 227 | Specification.read_spec bn_fun_strs bn_eq_strs lthy | |
| 228 | ||
| 229 | val bn_funs' = map prep_bn_fun bn_funs | |
| 230 | in | |
| 231 | lthy | |
| 232 | |> Local_Theory.theory (Sign.add_consts_i bn_funs') | |
| 233 | |> pair (bn_funs', bn_eqs) | |
| 234 | 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 | 235 | *} | 
| 
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 | 236 | |
| 
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 | 237 | ML {*
 | 
| 1284 | 238 | fun forth (_, _, _, x) = x | 
| 239 | ||
| 1285 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 240 | fun find_all eq xs k' = | 
| 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 241 | maps (fn (k, v) => if eq (k, k') then [v] else []) xs | 
| 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 242 | *} | 
| 1284 | 243 | |
| 1285 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 244 | 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 | 245 | 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 | 246 | 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 | 247 | fun mapp (_: int) [] = [] | 
| 1285 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 248 | | 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 | 249 | 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 | 250 | 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 | 251 | | 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 | 252 | in mapp 0 xs end | 
| 1285 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 253 | *} | 
| 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 | |
| 1285 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 255 | 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 | 256 | 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 | 257 | 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 | 258 | SOME x => 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 | 259 |   | NONE => error ("cannot find " ^ x ^ " in the binding specification.")  
 | 
| 
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 | *} | 
| 
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 | |
| 978 
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
973diff
changeset | 262 | ML {*
 | 
| 1284 | 263 | 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 | 264 | let | 
| 1285 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 265 | fun extract_cnstrs dt_strs = | 
| 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 266 | map ((map (fn (_, antys, _, bns) => (antys, bns))) o forth) dt_strs | 
| 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 267 | |
| 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 268 | fun prep_bn env bn_str = | 
| 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 269 | case (Syntax.read_term lthy bn_str) of | 
| 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 270 | Free (x, _) => (env_lookup env x, NONE) | 
| 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 271 | | Const (a, T) $ Free (x, _) => (env_lookup env x, SOME (Const (a, T))) | 
| 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 272 | | _ => 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 | 273 | |
| 1284 | 274 | fun prep_typ env (opt_name, _) = | 
| 1285 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 275 | case opt_name of | 
| 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 276 | NONE => [] | 
| 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 277 | | SOME x => find_all (op=) env x; | 
| 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 | 278 | |
| 1285 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 279 | fun prep_binds (anno_tys, bind_strs) = | 
| 1284 | 280 | let | 
| 1285 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 281 | val env = mk_env (map fst anno_tys) | 
| 1284 | 282 | val binds = map (fn (x, y) => (x, prep_bn env y)) bind_strs | 
| 283 | in | |
| 284 | map (prep_typ binds) anno_tys | |
| 285 | end | |
| 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 | 286 | in | 
| 1285 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 287 | map (map prep_binds) (extract_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 | 288 | end | 
| 966 
8ba35ce16f7e
Some cleaning of thy vs lthy vs context.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
962diff
changeset | 289 | *} | 
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 290 | |
| 1284 | 291 | ML {*
 | 
| 292 | fun nominal_datatype2_cmd (dt_strs, bn_fun_strs, bn_eq_strs) lthy = | |
| 293 | let | |
| 294 | fun prep_typ (tvs, tname, mx, _) = (tname, length tvs, mx) | |
| 295 | ||
| 296 | val ((dts, (bn_fun, bn_eq)), binds) = | |
| 297 | lthy | |
| 298 | |> Local_Theory.theory (Sign.add_types (map prep_typ dt_strs)) | |
| 299 | |> prepare_dts dt_strs | |
| 300 | ||>> prepare_bn_funs bn_fun_strs bn_eq_strs | |
| 301 | ||> prepare_binds dt_strs | |
| 302 | ||
| 303 | val _ = tracing (PolyML.makestring binds) | |
| 304 | in | |
| 305 | nominal_datatype2 dts bn_fun bn_eq binds lthy |> snd | |
| 306 | end | |
| 307 | *} | |
| 308 | ||
| 309 | ||
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 310 | (* 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 | 311 | |
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 312 | ML {*
 | 
| 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 313 | let | 
| 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 314 | 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 | 315 | 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 | 316 | 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 | 317 | (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 | 318 | end | 
| 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 319 | *} | 
| 
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 | |
| 961 
0f88e04eb486
Variable takes a 'name'.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
954diff
changeset | 322 | end | 
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 323 | |
| 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 324 | |
| 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 | 325 |