| author | Cezary Kaliszyk <kaliszyk@in.tum.de> | 
| Thu, 18 Mar 2010 08:03:42 +0100 | |
| changeset 1497 | 1c9931e5039a | 
| parent 1496 | fd483d8f486b | 
| child 1510 | be911e869fde | 
| 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 | 
| 1494 
923413256cbb
Clean 'Lift', start working only on exported things in Parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1486diff
changeset | 2 | imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" "Perm" "Fv" "Rsp" "Lift" | 
| 954 
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 {* 
 | 
| 1293 
dca51a1f0c0d
rawified the bind specs (ready to be used now)
 Christian Urban <urbanc@in.tum.de> parents: 
1290diff
changeset | 147 | fun rawify_binds dts_env cnstrs_env bn_fun_env binds = | 
| 
dca51a1f0c0d
rawified the bind specs (ready to be used now)
 Christian Urban <urbanc@in.tum.de> parents: 
1290diff
changeset | 148 | map (map (map (map (fn (opt_trm, i, j) => | 
| 1375 
aa787c9b6955
A version of Fv that takes into account recursive and non-recursive bindings.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1364diff
changeset | 149 | (Option.map (apfst (replace_term (cnstrs_env @ bn_fun_env) dts_env)) opt_trm, i, j))))) binds | 
| 1293 
dca51a1f0c0d
rawified the bind specs (ready to be used now)
 Christian Urban <urbanc@in.tum.de> parents: 
1290diff
changeset | 150 | *} | 
| 
dca51a1f0c0d
rawified the bind specs (ready to be used now)
 Christian Urban <urbanc@in.tum.de> parents: 
1290diff
changeset | 151 | |
| 1380 
dab8d99b37c1
added bn-information, but it is not yet ordered according to the dts
 Christian Urban <urbanc@in.tum.de> parents: 
1375diff
changeset | 152 | ML {*
 | 
| 1392 
baa67b07f436
parser produces ordered bn-fun information
 Christian Urban <urbanc@in.tum.de> parents: 
1380diff
changeset | 153 | fun find [] _ = error ("cannot find element")
 | 
| 
baa67b07f436
parser produces ordered bn-fun information
 Christian Urban <urbanc@in.tum.de> parents: 
1380diff
changeset | 154 | | find ((x, z)::xs) y = if (Long_Name.base_name x) = y then z else find xs y | 
| 
baa67b07f436
parser produces ordered bn-fun information
 Christian Urban <urbanc@in.tum.de> parents: 
1380diff
changeset | 155 | *} | 
| 
baa67b07f436
parser produces ordered bn-fun information
 Christian Urban <urbanc@in.tum.de> parents: 
1380diff
changeset | 156 | |
| 
baa67b07f436
parser produces ordered bn-fun information
 Christian Urban <urbanc@in.tum.de> parents: 
1380diff
changeset | 157 | ML {*
 | 
| 1403 
4a10338c2535
almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
 Christian Urban <urbanc@in.tum.de> parents: 
1392diff
changeset | 158 | fun prep_bn dt_names dts eqs = | 
| 1380 
dab8d99b37c1
added bn-information, but it is not yet ordered according to the dts
 Christian Urban <urbanc@in.tum.de> parents: 
1375diff
changeset | 159 | let | 
| 
dab8d99b37c1
added bn-information, but it is not yet ordered according to the dts
 Christian Urban <urbanc@in.tum.de> parents: 
1375diff
changeset | 160 | fun aux eq = | 
| 
dab8d99b37c1
added bn-information, but it is not yet ordered according to the dts
 Christian Urban <urbanc@in.tum.de> parents: 
1375diff
changeset | 161 | let | 
| 
dab8d99b37c1
added bn-information, but it is not yet ordered according to the dts
 Christian Urban <urbanc@in.tum.de> parents: 
1375diff
changeset | 162 | val (lhs, rhs) = eq | 
| 
dab8d99b37c1
added bn-information, but it is not yet ordered according to the dts
 Christian Urban <urbanc@in.tum.de> parents: 
1375diff
changeset | 163 | |> strip_qnt_body "all" | 
| 
dab8d99b37c1
added bn-information, but it is not yet ordered according to the dts
 Christian Urban <urbanc@in.tum.de> parents: 
1375diff
changeset | 164 | |> HOLogic.dest_Trueprop | 
| 
dab8d99b37c1
added bn-information, but it is not yet ordered according to the dts
 Christian Urban <urbanc@in.tum.de> parents: 
1375diff
changeset | 165 | |> HOLogic.dest_eq | 
| 1392 
baa67b07f436
parser produces ordered bn-fun information
 Christian Urban <urbanc@in.tum.de> parents: 
1380diff
changeset | 166 | val (bn_fun, [cnstr]) = strip_comb lhs | 
| 
baa67b07f436
parser produces ordered bn-fun information
 Christian Urban <urbanc@in.tum.de> parents: 
1380diff
changeset | 167 | val (_, ty) = dest_Free bn_fun | 
| 1380 
dab8d99b37c1
added bn-information, but it is not yet ordered according to the dts
 Christian Urban <urbanc@in.tum.de> parents: 
1375diff
changeset | 168 | val (ty_name, _) = dest_Type (domain_type ty) | 
| 
dab8d99b37c1
added bn-information, but it is not yet ordered according to the dts
 Christian Urban <urbanc@in.tum.de> parents: 
1375diff
changeset | 169 | val dt_index = find_index (fn x => x = ty_name) dt_names | 
| 1392 
baa67b07f436
parser produces ordered bn-fun information
 Christian Urban <urbanc@in.tum.de> parents: 
1380diff
changeset | 170 | val (cnstr_head, cnstr_args) = strip_comb cnstr | 
| 1380 
dab8d99b37c1
added bn-information, but it is not yet ordered according to the dts
 Christian Urban <urbanc@in.tum.de> parents: 
1375diff
changeset | 171 | val included = map (fn i => length (cnstr_args) - i - 1) (loose_bnos rhs) | 
| 
dab8d99b37c1
added bn-information, but it is not yet ordered according to the dts
 Christian Urban <urbanc@in.tum.de> parents: 
1375diff
changeset | 172 | in | 
| 1392 
baa67b07f436
parser produces ordered bn-fun information
 Christian Urban <urbanc@in.tum.de> parents: 
1380diff
changeset | 173 | (dt_index, (bn_fun, (cnstr_head, included))) | 
| 
baa67b07f436
parser produces ordered bn-fun information
 Christian Urban <urbanc@in.tum.de> parents: 
1380diff
changeset | 174 | end | 
| 
baa67b07f436
parser produces ordered bn-fun information
 Christian Urban <urbanc@in.tum.de> parents: 
1380diff
changeset | 175 | |
| 
baa67b07f436
parser produces ordered bn-fun information
 Christian Urban <urbanc@in.tum.de> parents: 
1380diff
changeset | 176 | fun order dts i ts = | 
| 
baa67b07f436
parser produces ordered bn-fun information
 Christian Urban <urbanc@in.tum.de> parents: 
1380diff
changeset | 177 | let | 
| 
baa67b07f436
parser produces ordered bn-fun information
 Christian Urban <urbanc@in.tum.de> parents: 
1380diff
changeset | 178 | val dt = nth dts i | 
| 
baa67b07f436
parser produces ordered bn-fun information
 Christian Urban <urbanc@in.tum.de> parents: 
1380diff
changeset | 179 | val cts = map (fn (x, _, _) => Binding.name_of x) ((fn (_, _, _, x) => x) dt) | 
| 
baa67b07f436
parser produces ordered bn-fun information
 Christian Urban <urbanc@in.tum.de> parents: 
1380diff
changeset | 180 | val ts' = map (fn (x, y) => (fst (dest_Const x), y)) ts | 
| 
baa67b07f436
parser produces ordered bn-fun information
 Christian Urban <urbanc@in.tum.de> parents: 
1380diff
changeset | 181 | in | 
| 
baa67b07f436
parser produces ordered bn-fun information
 Christian Urban <urbanc@in.tum.de> parents: 
1380diff
changeset | 182 | map (find ts') cts | 
| 
baa67b07f436
parser produces ordered bn-fun information
 Christian Urban <urbanc@in.tum.de> parents: 
1380diff
changeset | 183 | end | 
| 
baa67b07f436
parser produces ordered bn-fun information
 Christian Urban <urbanc@in.tum.de> parents: 
1380diff
changeset | 184 | |
| 
baa67b07f436
parser produces ordered bn-fun information
 Christian Urban <urbanc@in.tum.de> parents: 
1380diff
changeset | 185 | val unordered = AList.group (op=) (map aux eqs) | 
| 
baa67b07f436
parser produces ordered bn-fun information
 Christian Urban <urbanc@in.tum.de> parents: 
1380diff
changeset | 186 | val unordered' = map (fn (x, y) => (x, AList.group (op=) y)) unordered | 
| 
baa67b07f436
parser produces ordered bn-fun information
 Christian Urban <urbanc@in.tum.de> parents: 
1380diff
changeset | 187 | val ordered = map (fn (x, y) => (x, map (fn (v, z) => (v, order dts x z)) y)) unordered' | 
| 1380 
dab8d99b37c1
added bn-information, but it is not yet ordered according to the dts
 Christian Urban <urbanc@in.tum.de> parents: 
1375diff
changeset | 188 | in | 
| 1392 
baa67b07f436
parser produces ordered bn-fun information
 Christian Urban <urbanc@in.tum.de> parents: 
1380diff
changeset | 189 | ordered | 
| 1380 
dab8d99b37c1
added bn-information, but it is not yet ordered according to the dts
 Christian Urban <urbanc@in.tum.de> parents: 
1375diff
changeset | 190 | end | 
| 
dab8d99b37c1
added bn-information, but it is not yet ordered according to the dts
 Christian Urban <urbanc@in.tum.de> parents: 
1375diff
changeset | 191 | *} | 
| 
dab8d99b37c1
added bn-information, but it is not yet ordered according to the dts
 Christian Urban <urbanc@in.tum.de> parents: 
1375diff
changeset | 192 | |
| 1293 
dca51a1f0c0d
rawified the bind specs (ready to be used now)
 Christian Urban <urbanc@in.tum.de> parents: 
1290diff
changeset | 193 | ML {* 
 | 
| 1289 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 194 | fun add_primrec_wrapper funs eqs lthy = | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 195 | if null funs then (([], []), lthy) | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 196 | else | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 197 | let | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 198 | 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 | 199 | 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 | 200 | in | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 201 | Primrec.add_primrec funs' eqs' lthy | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 202 | end | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 203 | *} | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 204 | |
| 1293 
dca51a1f0c0d
rawified the bind specs (ready to be used now)
 Christian Urban <urbanc@in.tum.de> parents: 
1290diff
changeset | 205 | ML {*
 | 
| 
dca51a1f0c0d
rawified the bind specs (ready to be used now)
 Christian Urban <urbanc@in.tum.de> parents: 
1290diff
changeset | 206 | fun add_datatype_wrapper dt_names dts = | 
| 
dca51a1f0c0d
rawified the bind specs (ready to be used now)
 Christian Urban <urbanc@in.tum.de> parents: 
1290diff
changeset | 207 | let | 
| 
dca51a1f0c0d
rawified the bind specs (ready to be used now)
 Christian Urban <urbanc@in.tum.de> parents: 
1290diff
changeset | 208 | val conf = Datatype.default_config | 
| 
dca51a1f0c0d
rawified the bind specs (ready to be used now)
 Christian Urban <urbanc@in.tum.de> parents: 
1290diff
changeset | 209 | in | 
| 
dca51a1f0c0d
rawified the bind specs (ready to be used now)
 Christian Urban <urbanc@in.tum.de> parents: 
1290diff
changeset | 210 | Local_Theory.theory_result (Datatype.add_datatype conf dt_names dts) | 
| 
dca51a1f0c0d
rawified the bind specs (ready to be used now)
 Christian Urban <urbanc@in.tum.de> parents: 
1290diff
changeset | 211 | end | 
| 
dca51a1f0c0d
rawified the bind specs (ready to be used now)
 Christian Urban <urbanc@in.tum.de> parents: 
1290diff
changeset | 212 | *} | 
| 1289 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 213 | |
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 214 | ML {* 
 | 
| 1293 
dca51a1f0c0d
rawified the bind specs (ready to be used now)
 Christian Urban <urbanc@in.tum.de> parents: 
1290diff
changeset | 215 | fun raw_nominal_decls 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 | 216 | 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 | 217 | val thy = ProofContext.theory_of lthy | 
| 1289 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 218 | val thy_name = Context.theory_name thy | 
| 1251 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 219 | |
| 1289 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 220 | 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 | 221 | 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 | 222 | 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 | 223 | 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 | 224 | |
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 225 | val cnstrs = get_cnstr_strs dts | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 226 | val cnstrs_ty = get_typed_cnstrs dts | 
| 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 227 | 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 | 228 | 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 | 229 | (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 | 230 | 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 | 231 | |
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 232 | 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 | 233 | val bn_fun_strs' = add_raws bn_fun_strs | 
| 1266 | 234 | val bn_fun_env = bn_fun_strs ~~ bn_fun_strs' | 
| 1293 
dca51a1f0c0d
rawified the bind specs (ready to be used now)
 Christian Urban <urbanc@in.tum.de> parents: 
1290diff
changeset | 235 | val bn_fun_full_env = map (pairself (Long_Name.qualify thy_name)) | 
| 
dca51a1f0c0d
rawified the bind specs (ready to be used now)
 Christian Urban <urbanc@in.tum.de> parents: 
1290diff
changeset | 236 | (bn_fun_strs ~~ bn_fun_strs') | 
| 
dca51a1f0c0d
rawified the bind specs (ready to be used now)
 Christian Urban <urbanc@in.tum.de> parents: 
1290diff
changeset | 237 | |
| 1289 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 238 | 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 | 239 | |
| 1289 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 240 | val (raw_bn_funs, raw_bn_eqs) = rawify_bn_funs dts_env cnstrs_env bn_fun_env bn_funs bn_eqs | 
| 1293 
dca51a1f0c0d
rawified the bind specs (ready to be used now)
 Christian Urban <urbanc@in.tum.de> parents: 
1290diff
changeset | 241 | |
| 
dca51a1f0c0d
rawified the bind specs (ready to be used now)
 Christian Urban <urbanc@in.tum.de> parents: 
1290diff
changeset | 242 | val raw_binds = rawify_binds dts_env cnstrs_env bn_fun_full_env binds | 
| 1380 
dab8d99b37c1
added bn-information, but it is not yet ordered according to the dts
 Christian Urban <urbanc@in.tum.de> parents: 
1375diff
changeset | 243 | |
| 1403 
4a10338c2535
almost done with showing the equivalence between old and new alpha-equivalence (one subgoal remaining)
 Christian Urban <urbanc@in.tum.de> parents: 
1392diff
changeset | 244 | val raw_bns = prep_bn dt_full_names' raw_dts (map snd raw_bn_eqs) | 
| 1380 
dab8d99b37c1
added bn-information, but it is not yet ordered according to the dts
 Christian Urban <urbanc@in.tum.de> parents: 
1375diff
changeset | 245 | |
| 1480 
21cbb5b0e321
cheat_alpha_eqvt no longer needed. Cleaned the tracing messages.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1467diff
changeset | 246 | (*val _ = tracing (cat_lines (map PolyML.makestring raw_bns))*) | 
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 247 | in | 
| 1289 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 248 | lthy | 
| 1293 
dca51a1f0c0d
rawified the bind specs (ready to be used now)
 Christian Urban <urbanc@in.tum.de> parents: 
1290diff
changeset | 249 | |> add_datatype_wrapper raw_dt_names raw_dts | 
| 1289 
02eb0f600630
further code-refactoring in the parser
 Christian Urban <urbanc@in.tum.de> parents: 
1287diff
changeset | 250 | ||>> add_primrec_wrapper raw_bn_funs raw_bn_eqs | 
| 1293 
dca51a1f0c0d
rawified the bind specs (ready to be used now)
 Christian Urban <urbanc@in.tum.de> parents: 
1290diff
changeset | 251 | ||>> pair raw_binds | 
| 1380 
dab8d99b37c1
added bn-information, but it is not yet ordered according to the dts
 Christian Urban <urbanc@in.tum.de> parents: 
1375diff
changeset | 252 | ||>> pair raw_bns | 
| 1293 
dca51a1f0c0d
rawified the bind specs (ready to be used now)
 Christian Urban <urbanc@in.tum.de> parents: 
1290diff
changeset | 253 | end | 
| 
dca51a1f0c0d
rawified the bind specs (ready to be used now)
 Christian Urban <urbanc@in.tum.de> parents: 
1290diff
changeset | 254 | *} | 
| 
dca51a1f0c0d
rawified the bind specs (ready to be used now)
 Christian Urban <urbanc@in.tum.de> parents: 
1290diff
changeset | 255 | |
| 1341 
c25f797c7e6e
Added lifting of pseudo-injectivity, commented out the code again and enabled the weird examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1340diff
changeset | 256 | lemma equivp_hack: "equivp x" | 
| 
c25f797c7e6e
Added lifting of pseudo-injectivity, commented out the code again and enabled the weird examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1340diff
changeset | 257 | sorry | 
| 
c25f797c7e6e
Added lifting of pseudo-injectivity, commented out the code again and enabled the weird examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1340diff
changeset | 258 | ML {*
 | 
| 
c25f797c7e6e
Added lifting of pseudo-injectivity, commented out the code again and enabled the weird examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1340diff
changeset | 259 | fun equivp_hack ctxt rel = | 
| 
c25f797c7e6e
Added lifting of pseudo-injectivity, commented out the code again and enabled the weird examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1340diff
changeset | 260 | let | 
| 
c25f797c7e6e
Added lifting of pseudo-injectivity, commented out the code again and enabled the weird examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1340diff
changeset | 261 | val thy = ProofContext.theory_of ctxt | 
| 
c25f797c7e6e
Added lifting of pseudo-injectivity, commented out the code again and enabled the weird examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1340diff
changeset | 262 | val ty = domain_type (fastype_of rel) | 
| 
c25f797c7e6e
Added lifting of pseudo-injectivity, commented out the code again and enabled the weird examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1340diff
changeset | 263 | val cty = ctyp_of thy ty | 
| 
c25f797c7e6e
Added lifting of pseudo-injectivity, commented out the code again and enabled the weird examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1340diff
changeset | 264 | val ct = cterm_of thy rel | 
| 
c25f797c7e6e
Added lifting of pseudo-injectivity, commented out the code again and enabled the weird examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1340diff
changeset | 265 | in | 
| 
c25f797c7e6e
Added lifting of pseudo-injectivity, commented out the code again and enabled the weird examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1340diff
changeset | 266 |   Drule.instantiate' [SOME cty] [SOME ct] @{thm equivp_hack}
 | 
| 
c25f797c7e6e
Added lifting of pseudo-injectivity, commented out the code again and enabled the weird examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1340diff
changeset | 267 | end | 
| 
c25f797c7e6e
Added lifting of pseudo-injectivity, commented out the code again and enabled the weird examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1340diff
changeset | 268 | *} | 
| 
c25f797c7e6e
Added lifting of pseudo-injectivity, commented out the code again and enabled the weird examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1340diff
changeset | 269 | |
| 1420 
e655ea5f0b5f
The cheats described explicitely.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1418diff
changeset | 270 | (* These 2 are critical, we don't know how to do it in term5 *) | 
| 1484 | 271 | ML {* val cheat_fv_rsp = Unsynchronized.ref true *}
 | 
| 272 | ML {* val cheat_const_rsp = Unsynchronized.ref true *} (* For alpha_bn 0 and alpha_bn_rsp *)
 | |
| 1416 | 273 | |
| 1484 | 274 | ML {* val cheat_equivp = Unsynchronized.ref true *}
 | 
| 1443 
d78c093aebeb
explicit flag "cheat_equivp"
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1436diff
changeset | 275 | |
| 1420 
e655ea5f0b5f
The cheats described explicitely.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1418diff
changeset | 276 | (* Fixes for these 2 are known *) | 
| 1484 | 277 | ML {* val cheat_fv_eqvt = Unsynchronized.ref true *} (* The tactic works, building the goal needs fixing *)
 | 
| 278 | ML {* val cheat_alpha_eqvt = Unsynchronized.ref true *} (* The tactic works, building the goal needs fixing *)
 | |
| 1341 
c25f797c7e6e
Added lifting of pseudo-injectivity, commented out the code again and enabled the weird examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1340diff
changeset | 279 | |
| 1428 
4029105011ca
Show that the new types are in finite support typeclass.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1422diff
changeset | 280 | |
| 1346 
998b1bde64e7
Lift fv and bn eqvts; no need to lift alpha_eqvt.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1342diff
changeset | 281 | ML {*
 | 
| 1293 
dca51a1f0c0d
rawified the bind specs (ready to be used now)
 Christian Urban <urbanc@in.tum.de> parents: 
1290diff
changeset | 282 | fun nominal_datatype2 dts bn_funs bn_eqs binds lthy = | 
| 
dca51a1f0c0d
rawified the bind specs (ready to be used now)
 Christian Urban <urbanc@in.tum.de> parents: 
1290diff
changeset | 283 | let | 
| 1480 
21cbb5b0e321
cheat_alpha_eqvt no longer needed. Cleaned the tracing messages.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1467diff
changeset | 284 | val _ = tracing "Raw declarations"; | 
| 1340 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1339diff
changeset | 285 | val thy = ProofContext.theory_of lthy | 
| 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1339diff
changeset | 286 | val thy_name = Context.theory_name thy | 
| 1380 
dab8d99b37c1
added bn-information, but it is not yet ordered according to the dts
 Christian Urban <urbanc@in.tum.de> parents: 
1375diff
changeset | 287 | val ((((raw_dt_names, (raw_bn_funs_loc, raw_bn_eqs_loc)), raw_binds), raw_bns), lthy2) = | 
| 1293 
dca51a1f0c0d
rawified the bind specs (ready to be used now)
 Christian Urban <urbanc@in.tum.de> parents: 
1290diff
changeset | 288 | raw_nominal_decls dts bn_funs bn_eqs binds lthy | 
| 1396 
08294f4d6c08
Linked parser to fv and alpha.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1392diff
changeset | 289 | val morphism_2_1 = ProofContext.export_morphism lthy2 lthy | 
| 
08294f4d6c08
Linked parser to fv and alpha.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1392diff
changeset | 290 | val raw_bns_exp = map (apsnd (map (apfst (Morphism.term morphism_2_1)))) raw_bns; | 
| 
08294f4d6c08
Linked parser to fv and alpha.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1392diff
changeset | 291 | val bn_funs_decls = flat (map (fn (ith, l) => map (fn (bn, data) => (bn, ith, data)) l) raw_bns_exp); | 
| 1316 
0577afdb1732
Porting from Lift to Parser; until defining the Quotient type.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1309diff
changeset | 292 | val raw_bn_funs = map (Morphism.term morphism_2_1) raw_bn_funs_loc | 
| 
0577afdb1732
Porting from Lift to Parser; until defining the Quotient type.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1309diff
changeset | 293 | val raw_bn_eqs = ProofContext.export lthy2 lthy raw_bn_eqs_loc | 
| 1497 
1c9931e5039a
Move most of the exporting out of the parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1496diff
changeset | 294 | |
| 1295 
0ecc775e5fce
Link calls to Raw permutations, FV definition and alpha_definition into the parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1293diff
changeset | 295 | val dtinfo = Datatype.the_info (ProofContext.theory_of lthy2) (hd raw_dt_names); | 
| 1309 | 296 | val descr = #descr dtinfo; | 
| 1316 
0577afdb1732
Porting from Lift to Parser; until defining the Quotient type.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1309diff
changeset | 297 | val sorts = #sorts dtinfo; | 
| 
0577afdb1732
Porting from Lift to Parser; until defining the Quotient type.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1309diff
changeset | 298 | val all_typs = map (fn i => typ_of_dtyp descr sorts (DtRec i)) (map fst descr) | 
| 1309 | 299 | val all_full_tnames = map (fn (_, (n, _, _)) => n) descr; | 
| 300 | val dtinfos = map (Datatype.the_info (ProofContext.theory_of lthy2)) all_full_tnames; | |
| 1342 | 301 | val rel_dtinfos = List.take (dtinfos, (length dts)); | 
| 1309 | 302 | val inject = flat (map #inject dtinfos); | 
| 1428 
4029105011ca
Show that the new types are in finite support typeclass.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1422diff
changeset | 303 | val distincts = flat (map #distinct dtinfos); | 
| 1342 | 304 | val rel_distinct = map #distinct rel_dtinfos; | 
| 1316 
0577afdb1732
Porting from Lift to Parser; until defining the Quotient type.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1309diff
changeset | 305 | val induct = #induct dtinfo; | 
| 1309 | 306 | val inducts = #inducts dtinfo; | 
| 1480 
21cbb5b0e321
cheat_alpha_eqvt no longer needed. Cleaned the tracing messages.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1467diff
changeset | 307 | val _ = tracing "Defining permutations, fv and alpha"; | 
| 1295 
0ecc775e5fce
Link calls to Raw permutations, FV definition and alpha_definition into the parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1293diff
changeset | 308 | val ((raw_perm_def, raw_perm_simps, perms), lthy3) = | 
| 
0ecc775e5fce
Link calls to Raw permutations, FV definition and alpha_definition into the parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1293diff
changeset | 309 | Local_Theory.theory_result (define_raw_perms dtinfo (length dts)) lthy2; | 
| 
0ecc775e5fce
Link calls to Raw permutations, FV definition and alpha_definition into the parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1293diff
changeset | 310 | val raw_binds_flat = map (map flat) raw_binds; | 
| 1497 
1c9931e5039a
Move most of the exporting out of the parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1496diff
changeset | 311 | val (((fv_ts, fv_def), ((alpha_ts, alpha_intros), (alpha_cases, alpha_induct))), lthy4) = | 
| 
1c9931e5039a
Move most of the exporting out of the parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1496diff
changeset | 312 | define_fv_alpha_export dtinfo raw_binds_flat bn_funs_decls lthy3; | 
| 1448 
f2c50884dfb9
fv_eqvt_cheat no longer needed.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1447diff
changeset | 313 | val (fv_ts_nobn, fv_ts_bn) = chop (length perms) fv_ts; | 
| 1416 | 314 | val (alpha_ts_nobn, alpha_ts_bn) = chop (length perms) alpha_ts | 
| 1340 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1339diff
changeset | 315 | val alpha_inducts = Project_Rule.projects lthy4 (1 upto (length dts)) alpha_induct | 
| 1309 | 316 | val dts_names = map (fn (i, (s, _, _)) => (s, i)) (#descr dtinfo); | 
| 317 | val bn_tys = map (domain_type o fastype_of) raw_bn_funs; | |
| 318 | val bn_nos = map (dtyp_no_of_typ dts_names) bn_tys; | |
| 1497 
1c9931e5039a
Move most of the exporting out of the parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1496diff
changeset | 319 | val bns = raw_bn_funs ~~ bn_nos; | 
| 1496 
fd483d8f486b
Prove pseudo-inject (eq-iff) on the exported level and rename appropriately.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1495diff
changeset | 320 | val alpha_eq_iff = build_alpha_inj alpha_intros (inject @ distincts) alpha_cases lthy4 | 
| 1480 
21cbb5b0e321
cheat_alpha_eqvt no longer needed. Cleaned the tracing messages.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1467diff
changeset | 321 | val _ = tracing "Proving equivariance"; | 
| 1409 
25b02cc185e2
build_eqvts no longer requires permutations.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1408diff
changeset | 322 | val (bv_eqvts, lthy5) = fold_map (build_bv_eqvt (raw_bn_eqs @ raw_perm_def) inducts) bns lthy4; | 
| 1411 
6b9833936281
explicit cheat_fv_eqvt
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1410diff
changeset | 323 | val fv_eqvt_tac = | 
| 
6b9833936281
explicit cheat_fv_eqvt
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1410diff
changeset | 324 | if !cheat_fv_eqvt then (fn _ => fn _ => Skip_Proof.cheat_tac thy) | 
| 1495 
219e3ff68de8
Prove eqvts on exported terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1494diff
changeset | 325 | else build_eqvts_tac induct ((flat (map snd bv_eqvts)) @ fv_def @ raw_perm_def) lthy5 | 
| 
219e3ff68de8
Prove eqvts on exported terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1494diff
changeset | 326 | val (fv_eqvts, lthy6) = build_eqvts Binding.empty fv_ts_nobn fv_eqvt_tac lthy5; | 
| 
219e3ff68de8
Prove eqvts on exported terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1494diff
changeset | 327 | val (fv_bn_eqvts, lthy6a) = | 
| 
219e3ff68de8
Prove eqvts on exported terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1494diff
changeset | 328 | if fv_ts_bn = [] then ([], lthy6) else | 
| 
219e3ff68de8
Prove eqvts on exported terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1494diff
changeset | 329 | fold_map (build_bv_eqvt ((flat (map snd bv_eqvts)) @ fv_def @ raw_perm_def) inducts) | 
| 
219e3ff68de8
Prove eqvts on exported terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1494diff
changeset | 330 | (fv_ts_bn ~~ (map snd bns)) lthy6; | 
| 
219e3ff68de8
Prove eqvts on exported terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1494diff
changeset | 331 | val raw_fv_bv_eqvt = flat (map snd bv_eqvts) @ (snd fv_eqvts) @ flat (map snd fv_bn_eqvts) | 
| 1454 
7c8cd6eae8e2
FV_bn generated for recursive functions as well, and used in main fv for bindings.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1448diff
changeset | 332 | fun alpha_eqvt_tac' _ = | 
| 
7c8cd6eae8e2
FV_bn generated for recursive functions as well, and used in main fv for bindings.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1448diff
changeset | 333 | if !cheat_alpha_eqvt then Skip_Proof.cheat_tac thy | 
| 1496 
fd483d8f486b
Prove pseudo-inject (eq-iff) on the exported level and rename appropriately.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1495diff
changeset | 334 | else alpha_eqvt_tac alpha_induct (raw_perm_def @ alpha_eq_iff @ raw_fv_bv_eqvt) lthy6a 1 | 
| 1495 
219e3ff68de8
Prove eqvts on exported terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1494diff
changeset | 335 | val alpha_eqvt = build_alpha_eqvts alpha_ts alpha_eqvt_tac' lthy6a; | 
| 1480 
21cbb5b0e321
cheat_alpha_eqvt no longer needed. Cleaned the tracing messages.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1467diff
changeset | 336 | val _ = tracing "Proving equivalence"; | 
| 1494 
923413256cbb
Clean 'Lift', start working only on exported things in Parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1486diff
changeset | 337 | val alpha_equivp = | 
| 1495 
219e3ff68de8
Prove eqvts on exported terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1494diff
changeset | 338 | if !cheat_equivp then map (equivp_hack lthy6a) alpha_ts_nobn | 
| 1494 
923413256cbb
Clean 'Lift', start working only on exported things in Parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1486diff
changeset | 339 | else build_equivps alpha_ts induct alpha_induct | 
| 1496 
fd483d8f486b
Prove pseudo-inject (eq-iff) on the exported level and rename appropriately.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1495diff
changeset | 340 | inject alpha_eq_iff distincts alpha_cases alpha_eqvt lthy6a; | 
| 1412 
137cad9c1ce9
Equivp working only on the standard alpha-equivalences.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1411diff
changeset | 341 | val qty_binds = map (fn (_, b, _, _) => b) dts; | 
| 
137cad9c1ce9
Equivp working only on the standard alpha-equivalences.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1411diff
changeset | 342 | val qty_names = map Name.of_binding qty_binds; | 
| 
137cad9c1ce9
Equivp working only on the standard alpha-equivalences.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1411diff
changeset | 343 | val qty_full_names = map (Long_Name.qualify thy_name) qty_names | 
| 1316 
0577afdb1732
Porting from Lift to Parser; until defining the Quotient type.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1309diff
changeset | 344 | val lthy7 = define_quotient_type | 
| 1414 
d3b86738e848
Lifting constants works for all examples.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1412diff
changeset | 345 | (map (fn ((b, t), alpha) => (([], b, NoSyn), (t, alpha))) ((qty_binds ~~ all_typs) ~~ alpha_ts_nobn)) | 
| 1495 
219e3ff68de8
Prove eqvts on exported terms.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1494diff
changeset | 346 | (ALLGOALS (resolve_tac alpha_equivp)) lthy6a; | 
| 1340 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1339diff
changeset | 347 | val const_names = map Name.of_binding (flat (map (fn (_, _, _, t) => map (fn (b, _, _) => b) t) dts)); | 
| 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1339diff
changeset | 348 | val raw_consts = | 
| 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1339diff
changeset | 349 | flat (map (fn (i, (_, _, l)) => | 
| 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1339diff
changeset | 350 | map (fn (cname, dts) => | 
| 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1339diff
changeset | 351 | Const (cname, map (typ_of_dtyp descr sorts) dts ---> | 
| 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1339diff
changeset | 352 | typ_of_dtyp descr sorts (DtRec i))) l) descr); | 
| 1494 
923413256cbb
Clean 'Lift', start working only on exported things in Parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1486diff
changeset | 353 | val (consts, const_defs, lthy8) = quotient_lift_consts_export (const_names ~~ raw_consts) lthy7; | 
| 1428 
4029105011ca
Show that the new types are in finite support typeclass.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1422diff
changeset | 354 | (* Could be done nicer *) | 
| 
4029105011ca
Show that the new types are in finite support typeclass.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1422diff
changeset | 355 | val q_tys = distinct (op =) (map (snd o strip_type o fastype_of) consts); | 
| 1480 
21cbb5b0e321
cheat_alpha_eqvt no longer needed. Cleaned the tracing messages.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1467diff
changeset | 356 | val _ = tracing "Proving respects"; | 
| 1340 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1339diff
changeset | 357 | val (bns_rsp_pre, lthy9) = fold_map ( | 
| 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1339diff
changeset | 358 | fn (bn_t, i) => prove_const_rsp Binding.empty [bn_t] | 
| 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1339diff
changeset | 359 | (fn _ => fvbv_rsp_tac (nth alpha_inducts i) raw_bn_eqs 1)) bns lthy8; | 
| 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1339diff
changeset | 360 | val bns_rsp = flat (map snd bns_rsp_pre); | 
| 1416 | 361 | fun fv_rsp_tac _ = if !cheat_fv_rsp then Skip_Proof.cheat_tac thy | 
| 362 | else fvbv_rsp_tac alpha_induct fv_def 1; | |
| 363 | val ((_, fv_rsp), lthy10) = prove_const_rsp Binding.empty fv_ts fv_rsp_tac lthy9 | |
| 364 | val (perms_rsp, lthy11) = prove_const_rsp Binding.empty perms | |
| 365 | (fn _ => asm_simp_tac (HOL_ss addsimps alpha_eqvt) 1) lthy10; | |
| 366 | fun const_rsp_tac _ = if !cheat_const_rsp then Skip_Proof.cheat_tac thy | |
| 1496 
fd483d8f486b
Prove pseudo-inject (eq-iff) on the exported level and rename appropriately.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1495diff
changeset | 367 | else constr_rsp_tac alpha_eq_iff (fv_rsp @ bns_rsp) alpha_equivp 1 | 
| 1416 | 368 | val (const_rsps, lthy12) = fold_map (fn cnst => prove_const_rsp Binding.empty [cnst] | 
| 369 | const_rsp_tac) (raw_consts @ alpha_ts_bn) lthy11 | |
| 370 | val qfv_names = map (unsuffix "_raw" o Long_Name.base_name o fst o dest_Const) fv_ts | |
| 1340 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1339diff
changeset | 371 | val (qfv_defs, lthy12a) = fold_map Quotient_Def.quotient_lift_const (qfv_names ~~ fv_ts) lthy12; | 
| 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1339diff
changeset | 372 | val qbn_names = map (fn (b, _ , _) => Name.of_binding b) bn_funs | 
| 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1339diff
changeset | 373 | val (qbn_defs, lthy12b) = fold_map Quotient_Def.quotient_lift_const (qbn_names ~~ raw_bn_funs) lthy12a; | 
| 1417 
8f5f7abe22c1
Lift alpha_bn_constants.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1416diff
changeset | 374 | val qalpha_bn_names = map (unsuffix "_raw" o Long_Name.base_name o fst o dest_Const) alpha_ts_bn | 
| 1497 
1c9931e5039a
Move most of the exporting out of the parser.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1496diff
changeset | 375 | val (qbn_defs, lthy12c) = fold_map Quotient_Def.quotient_lift_const (qalpha_bn_names ~~ alpha_ts_bn) lthy12b; | 
| 1480 
21cbb5b0e321
cheat_alpha_eqvt no longer needed. Cleaned the tracing messages.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1467diff
changeset | 376 | val _ = tracing "Lifting permutations"; | 
| 1417 
8f5f7abe22c1
Lift alpha_bn_constants.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1416diff
changeset | 377 | val thy = Local_Theory.exit_global lthy12c; | 
| 1340 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1339diff
changeset | 378 | val perm_names = map (fn x => "permute_" ^ x) qty_names | 
| 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1339diff
changeset | 379 | val thy' = define_lifted_perms qty_full_names (perm_names ~~ perms) raw_perm_simps thy; | 
| 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1339diff
changeset | 380 | val lthy13 = Theory_Target.init NONE thy'; | 
| 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1339diff
changeset | 381 | val q_name = space_implode "_" qty_names; | 
| 1480 
21cbb5b0e321
cheat_alpha_eqvt no longer needed. Cleaned the tracing messages.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1467diff
changeset | 382 | val _ = tracing "Lifting induction"; | 
| 1422 
a26cb19375e8
Remove "_raw" from lifted theorems.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1420diff
changeset | 383 | val q_induct = lift_thm lthy13 induct; | 
| 1340 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1339diff
changeset | 384 | val (_, lthy14) = Local_Theory.note ((Binding.name (q_name ^ "_induct"), []), [q_induct]) lthy13; | 
| 1447 
378b8c791de8
derive "inducts" from "induct" instead of lifting again is much faster.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1445diff
changeset | 385 | val q_inducts = Project_Rule.projects lthy13 (1 upto (length alpha_inducts)) q_induct | 
| 
378b8c791de8
derive "inducts" from "induct" instead of lifting again is much faster.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1445diff
changeset | 386 | val (_, lthy14a) = Local_Theory.note ((Binding.name (q_name ^ "_inducts"), []), q_inducts) lthy14; | 
| 1422 
a26cb19375e8
Remove "_raw" from lifted theorems.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1420diff
changeset | 387 | val q_perm = map (lift_thm lthy14) raw_perm_def; | 
| 1447 
378b8c791de8
derive "inducts" from "induct" instead of lifting again is much faster.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1445diff
changeset | 388 | val (_, lthy15) = Local_Theory.note ((Binding.name (q_name ^ "_perm"), []), q_perm) lthy14a; | 
| 1422 
a26cb19375e8
Remove "_raw" from lifted theorems.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1420diff
changeset | 389 | val q_fv = map (lift_thm lthy15) fv_def; | 
| 1340 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1339diff
changeset | 390 | val (_, lthy16) = Local_Theory.note ((Binding.name (q_name ^ "_fv"), []), q_fv) lthy15; | 
| 1422 
a26cb19375e8
Remove "_raw" from lifted theorems.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1420diff
changeset | 391 | val q_bn = map (lift_thm lthy16) raw_bn_eqs; | 
| 1340 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1339diff
changeset | 392 | val (_, lthy17) = Local_Theory.note ((Binding.name (q_name ^ "_bn"), []), q_bn) lthy16; | 
| 1496 
fd483d8f486b
Prove pseudo-inject (eq-iff) on the exported level and rename appropriately.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1495diff
changeset | 393 | val _ = tracing "Lifting eq-iff"; | 
| 
fd483d8f486b
Prove pseudo-inject (eq-iff) on the exported level and rename appropriately.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1495diff
changeset | 394 |   val eq_iff_unfolded1 = map (Local_Defs.unfold lthy17 @{thms alpha_gen2}) alpha_eq_iff
 | 
| 
fd483d8f486b
Prove pseudo-inject (eq-iff) on the exported level and rename appropriately.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1495diff
changeset | 395 |   val eq_iff_unfolded2 = map (Local_Defs.unfold lthy17 @{thms alpha_gen}) eq_iff_unfolded1
 | 
| 
fd483d8f486b
Prove pseudo-inject (eq-iff) on the exported level and rename appropriately.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1495diff
changeset | 396 | val q_eq_iff_pre1 = map (lift_thm lthy17) eq_iff_unfolded2; | 
| 
fd483d8f486b
Prove pseudo-inject (eq-iff) on the exported level and rename appropriately.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1495diff
changeset | 397 |   val q_eq_iff_pre2 = map (Local_Defs.fold lthy17 @{thms alpha_gen2}) q_eq_iff_pre1
 | 
| 
fd483d8f486b
Prove pseudo-inject (eq-iff) on the exported level and rename appropriately.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1495diff
changeset | 398 |   val q_eq_iff = map (Local_Defs.fold lthy17 @{thms alpha_gen}) q_eq_iff_pre2
 | 
| 
fd483d8f486b
Prove pseudo-inject (eq-iff) on the exported level and rename appropriately.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1495diff
changeset | 399 | val (_, lthy18) = Local_Theory.note ((Binding.name (q_name ^ "_eq_iff"), []), q_eq_iff) lthy17; | 
| 1342 | 400 | val rel_dists = flat (map (distinct_rel lthy18 alpha_cases) | 
| 401 | (rel_distinct ~~ (List.take (alpha_ts, (length dts))))) | |
| 1422 
a26cb19375e8
Remove "_raw" from lifted theorems.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1420diff
changeset | 402 | val q_dis = map (lift_thm lthy18) rel_dists; | 
| 1346 
998b1bde64e7
Lift fv and bn eqvts; no need to lift alpha_eqvt.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1342diff
changeset | 403 | val (_, lthy19) = Local_Theory.note ((Binding.name (q_name ^ "_distinct"), []), q_dis) lthy18; | 
| 1422 
a26cb19375e8
Remove "_raw" from lifted theorems.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1420diff
changeset | 404 | val q_eqvt = map (lift_thm lthy19) raw_fv_bv_eqvt; | 
| 1346 
998b1bde64e7
Lift fv and bn eqvts; no need to lift alpha_eqvt.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1342diff
changeset | 405 | val (_, lthy20) = Local_Theory.note ((Binding.empty, | 
| 
998b1bde64e7
Lift fv and bn eqvts; no need to lift alpha_eqvt.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1342diff
changeset | 406 | [Attrib.internal (fn _ => Nominal_ThmDecls.eqvt_add)]), q_eqvt) lthy19; | 
| 1480 
21cbb5b0e321
cheat_alpha_eqvt no longer needed. Cleaned the tracing messages.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1467diff
changeset | 407 | val _ = tracing "Finite Support"; | 
| 1428 
4029105011ca
Show that the new types are in finite support typeclass.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1422diff
changeset | 408 | val supports = map (prove_supports lthy20 q_perm) consts | 
| 1433 
7a9217a7f681
Do not fail if the finite support proof fails.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1431diff
changeset | 409 | val fin_supp = HOLogic.conj_elims (prove_fs lthy20 q_induct supports q_tys) handle _ => [] | 
| 1428 
4029105011ca
Show that the new types are in finite support typeclass.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1422diff
changeset | 410 | val thy3 = Local_Theory.exit_global lthy20; | 
| 
4029105011ca
Show that the new types are in finite support typeclass.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1422diff
changeset | 411 |   val lthy21 = Theory_Target.instantiation (qty_full_names, [], @{sort fs}) thy3;
 | 
| 
4029105011ca
Show that the new types are in finite support typeclass.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1422diff
changeset | 412 | fun tac _ = Class.intro_classes_tac [] THEN (ALLGOALS (resolve_tac fin_supp)) | 
| 1433 
7a9217a7f681
Do not fail if the finite support proof fails.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1431diff
changeset | 413 | val lthy22 = Class.prove_instantiation_instance tac lthy21 handle _ => lthy20 | 
| 1293 
dca51a1f0c0d
rawified the bind specs (ready to be used now)
 Christian Urban <urbanc@in.tum.de> parents: 
1290diff
changeset | 414 | in | 
| 1428 
4029105011ca
Show that the new types are in finite support typeclass.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1422diff
changeset | 415 | ((raw_dt_names, raw_bn_funs, raw_bn_eqs, raw_binds), lthy22) | 
| 1346 
998b1bde64e7
Lift fv and bn eqvts; no need to lift alpha_eqvt.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1342diff
changeset | 416 | end | 
| 978 
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
973diff
changeset | 417 | *} | 
| 1340 
f201eb6acafc
Lift BV,FV,Permutations and injection :).
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1339diff
changeset | 418 | |
| 1422 
a26cb19375e8
Remove "_raw" from lifted theorems.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1420diff
changeset | 419 | |
| 1284 | 420 | ML {* 
 | 
| 421 | (* parsing the datatypes and declaring *) | |
| 422 | (* constructors in the local theory *) | |
| 423 | 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 | 424 | 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 | 425 | val thy = ProofContext.theory_of lthy | 
| 1284 | 426 | |
| 427 | fun mk_type full_tname tvrs = | |
| 428 | 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 | 429 | |
| 1284 | 430 | fun prep_cnstr lthy full_tname tvs (cname, anno_tys, mx, _) = | 
| 431 | let | |
| 432 | val tys = map (Syntax.read_typ lthy o snd) anno_tys | |
| 433 | val ty = mk_type full_tname tvs | |
| 434 | in | |
| 435 | ((cname, tys ---> ty, mx), (cname, tys, mx)) | |
| 436 | end | |
| 437 | ||
| 438 | fun prep_dt lthy (tvs, tname, mx, cnstrs) = | |
| 439 | let | |
| 440 | val full_tname = Sign.full_name thy tname | |
| 441 | val (cnstrs', cnstrs'') = | |
| 442 | split_list (map (prep_cnstr lthy full_tname tvs) cnstrs) | |
| 443 | in | |
| 444 | (cnstrs', (tvs, tname, mx, cnstrs'')) | |
| 445 | end | |
| 1266 | 446 | |
| 1284 | 447 | val (cnstrs, dts) = | 
| 448 | 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 | 449 | in | 
| 1284 | 450 | lthy | 
| 451 | |> Local_Theory.theory (Sign.add_consts_i (flat cnstrs)) | |
| 452 | |> 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 | 453 | end | 
| 
11b8798dea5d
parsing and definition of raw datatype and bv-function work (not very beautiful)
 Christian Urban <urbanc@in.tum.de> parents: 
1232diff
changeset | 454 | *} | 
| 1232 
35bc8f7b66f8
parsing of function definitions almost works now; still an error with undefined constants
 Christian Urban <urbanc@in.tum.de> parents: 
1228diff
changeset | 455 | |
| 1284 | 456 | ML {*
 | 
| 1285 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 457 | (* parsing the binding function specification and *) | 
| 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 458 | (* declaring the functions in the local theory *) | 
| 1284 | 459 | fun prepare_bn_funs bn_fun_strs bn_eq_strs lthy = | 
| 460 | let | |
| 461 | val ((bn_funs, bn_eqs), _) = | |
| 462 | Specification.read_spec bn_fun_strs bn_eq_strs lthy | |
| 463 | ||
| 1380 
dab8d99b37c1
added bn-information, but it is not yet ordered according to the dts
 Christian Urban <urbanc@in.tum.de> parents: 
1375diff
changeset | 464 | fun prep_bn_fun ((bn, T), mx) = (bn, T, mx) | 
| 1284 | 465 | val bn_funs' = map prep_bn_fun bn_funs | 
| 466 | in | |
| 467 | lthy | |
| 468 | |> Local_Theory.theory (Sign.add_consts_i bn_funs') | |
| 469 | |> pair (bn_funs', bn_eqs) | |
| 470 | 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 | 471 | *} | 
| 
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 | 472 | |
| 
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 | 473 | ML {*
 | 
| 1290 
a7e7ffb7f362
modified for new binding format - hope it is the intended one
 Christian Urban <urbanc@in.tum.de> parents: 
1289diff
changeset | 474 | 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 | 475 | 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 | 476 | *} | 
| 1284 | 477 | |
| 1285 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 478 | ML {*
 | 
| 1290 
a7e7ffb7f362
modified for new binding format - hope it is the intended one
 Christian Urban <urbanc@in.tum.de> parents: 
1289diff
changeset | 479 | (* 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 | 480 | 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 | 481 | 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 | 482 | fun mapp (_: int) [] = [] | 
| 1285 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 483 | | 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 | 484 | 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 | 485 | 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 | 486 | | 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 | 487 | in mapp 0 xs end | 
| 1285 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 488 | *} | 
| 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 | 489 | |
| 1285 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 490 | 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 | 491 | 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 | 492 | 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 | 493 | SOME x => x | 
| 1287 
8557af71724e
slight simplification of the raw-decl generation
 Christian Urban <urbanc@in.tum.de> parents: 
1285diff
changeset | 494 |   | 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 | 495 | *} | 
| 
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 | 496 | |
| 978 
b44592adf235
Improper interface for datatype and function packages and proper interface lateron.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
973diff
changeset | 497 | ML {*
 | 
| 1484 | 498 | val recursive = Unsynchronized.ref false | 
| 1436 
04dad9b0136d
started supp-fv proofs (is going to work)
 Christian Urban <urbanc@in.tum.de> parents: 
1434diff
changeset | 499 | *} | 
| 
04dad9b0136d
started supp-fv proofs (is going to work)
 Christian Urban <urbanc@in.tum.de> parents: 
1434diff
changeset | 500 | |
| 
04dad9b0136d
started supp-fv proofs (is going to work)
 Christian Urban <urbanc@in.tum.de> parents: 
1434diff
changeset | 501 | ML {*
 | 
| 1284 | 502 | 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 | 503 | let | 
| 1287 
8557af71724e
slight simplification of the raw-decl generation
 Christian Urban <urbanc@in.tum.de> parents: 
1285diff
changeset | 504 | 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 | 505 | 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 | 506 | |
| 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 507 | fun prep_bn env bn_str = | 
| 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 508 | 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 | 509 | Free (x, _) => (NONE, env_lookup env x) | 
| 1436 
04dad9b0136d
started supp-fv proofs (is going to work)
 Christian Urban <urbanc@in.tum.de> parents: 
1434diff
changeset | 510 | | Const (a, T) $ Free (x, _) => (SOME (Const (a, T), !recursive), env_lookup env x) | 
| 1285 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 511 | | _ => 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 | 512 | |
| 1290 
a7e7ffb7f362
modified for new binding format - hope it is the intended one
 Christian Urban <urbanc@in.tum.de> parents: 
1289diff
changeset | 513 | fun prep_typ env (i, opt_name) = | 
| 1285 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 514 | case opt_name of | 
| 
e3ac56d3b7af
added example from my phd
 Christian Urban <urbanc@in.tum.de> parents: 
1284diff
changeset | 515 | NONE => [] | 
| 1290 
a7e7ffb7f362
modified for new binding format - hope it is the intended one
 Christian Urban <urbanc@in.tum.de> parents: 
1289diff
changeset | 516 | | 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 | 517 | |
| 1287 
8557af71724e
slight simplification of the raw-decl generation
 Christian Urban <urbanc@in.tum.de> parents: 
1285diff
changeset | 518 | (* 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 | 519 | |
| 
8557af71724e
slight simplification of the raw-decl generation
 Christian Urban <urbanc@in.tum.de> parents: 
1285diff
changeset | 520 | fun prep_binds (annos, bind_strs) = | 
| 1284 | 521 | let | 
| 1290 
a7e7ffb7f362
modified for new binding format - hope it is the intended one
 Christian Urban <urbanc@in.tum.de> parents: 
1289diff
changeset | 522 | val env = mk_env annos (* for every label the index *) | 
| 1284 | 523 | val binds = map (fn (x, y) => (x, prep_bn env y)) bind_strs | 
| 524 | in | |
| 1290 
a7e7ffb7f362
modified for new binding format - hope it is the intended one
 Christian Urban <urbanc@in.tum.de> parents: 
1289diff
changeset | 525 | map_index (prep_typ binds) annos | 
| 1284 | 526 | end | 
| 1287 
8557af71724e
slight simplification of the raw-decl generation
 Christian Urban <urbanc@in.tum.de> parents: 
1285diff
changeset | 527 | |
| 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 | 528 | in | 
| 1290 
a7e7ffb7f362
modified for new binding format - hope it is the intended one
 Christian Urban <urbanc@in.tum.de> parents: 
1289diff
changeset | 529 | 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 | 530 | end | 
| 966 
8ba35ce16f7e
Some cleaning of thy vs lthy vs context.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
962diff
changeset | 531 | *} | 
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 532 | |
| 1284 | 533 | ML {*
 | 
| 534 | fun nominal_datatype2_cmd (dt_strs, bn_fun_strs, bn_eq_strs) lthy = | |
| 535 | let | |
| 536 | fun prep_typ (tvs, tname, mx, _) = (tname, length tvs, mx) | |
| 537 | ||
| 1380 
dab8d99b37c1
added bn-information, but it is not yet ordered according to the dts
 Christian Urban <urbanc@in.tum.de> parents: 
1375diff
changeset | 538 | val lthy0 = | 
| 
dab8d99b37c1
added bn-information, but it is not yet ordered according to the dts
 Christian Urban <urbanc@in.tum.de> parents: 
1375diff
changeset | 539 | Local_Theory.theory (Sign.add_types (map prep_typ dt_strs)) lthy | 
| 
dab8d99b37c1
added bn-information, but it is not yet ordered according to the dts
 Christian Urban <urbanc@in.tum.de> parents: 
1375diff
changeset | 540 | val (dts, lthy1) = | 
| 
dab8d99b37c1
added bn-information, but it is not yet ordered according to the dts
 Christian Urban <urbanc@in.tum.de> parents: 
1375diff
changeset | 541 | prepare_dts dt_strs lthy0 | 
| 
dab8d99b37c1
added bn-information, but it is not yet ordered according to the dts
 Christian Urban <urbanc@in.tum.de> parents: 
1375diff
changeset | 542 | val ((bn_funs, bn_eqs), lthy2) = | 
| 
dab8d99b37c1
added bn-information, but it is not yet ordered according to the dts
 Christian Urban <urbanc@in.tum.de> parents: 
1375diff
changeset | 543 | prepare_bn_funs bn_fun_strs bn_eq_strs lthy1 | 
| 
dab8d99b37c1
added bn-information, but it is not yet ordered according to the dts
 Christian Urban <urbanc@in.tum.de> parents: 
1375diff
changeset | 544 | val binds = prepare_binds dt_strs lthy2 | 
| 1284 | 545 | in | 
| 1380 
dab8d99b37c1
added bn-information, but it is not yet ordered according to the dts
 Christian Urban <urbanc@in.tum.de> parents: 
1375diff
changeset | 546 | nominal_datatype2 dts bn_funs bn_eqs binds lthy |> snd | 
| 1284 | 547 | end | 
| 548 | *} | |
| 549 | ||
| 550 | ||
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 551 | (* 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 | 552 | |
| 954 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 553 | ML {*
 | 
| 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 554 | let | 
| 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 555 | 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 | 556 | 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 | 557 | 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 | 558 | (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 | 559 | end | 
| 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 560 | *} | 
| 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 561 | |
| 
c009d2535896
very rough example file for how nominal2 specification can be parsed
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 562 | |
| 961 
0f88e04eb486
Variable takes a 'name'.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
954diff
changeset | 563 | end | 
| 1223 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 564 | |
| 
160343d86a6f
"raw"-ified the term-constructors and types given in the specification
 Christian Urban <urbanc@in.tum.de> parents: 
1194diff
changeset | 565 | |
| 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 | 566 |