Quot/Nominal/Test.thy
author Christian Urban <urbanc@in.tum.de>
Tue, 23 Feb 2010 16:31:40 +0100
changeset 1228 c179ad9d2446
parent 1223 160343d86a6f
child 1232 35bc8f7b66f8
permissions -rw-r--r--
declarartion of the raw datatype already works; raw binding functions throw an exception about mutual recursive types
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory Test
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
begin
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
atom_decl name
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
     8
(* tests *)
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
     9
ML {*
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    10
Datatype.datatype_cmd;
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    11
Datatype.add_datatype Datatype.default_config;
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    12
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    13
Primrec.add_primrec_cmd;
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    14
Primrec.add_primrec;
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    15
Primrec.add_primrec_simple;
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    16
*}
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    17
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    18
section {* test for setting up a primrec on the ML-level *}
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    19
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    20
section{* Interface for nominal_datatype *}
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    21
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    22
text {*
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    23
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    24
Nominal-Datatype-part:
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    25
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    26
1st Arg: string list  
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    27
         ^^^^^^^^^^^
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    28
         strings of the types to be defined
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    29
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    30
2nd Arg: (string list * binding * mixfix * (binding * typ list * mixfix) list) list
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    31
         ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^   ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    32
               type(s) to be defined             constructors list
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    33
               (ty args, name, syn)              (name, typs, syn)
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    34
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    35
Binder-Function-part:
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    36
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    37
3rd 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: 1194
diff changeset
    38
         ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^    
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    39
            binding function(s)           
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    40
              to be defined               
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    41
            (name, type, syn)             
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    42
1223
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
    43
4th Arg:  term list 
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
    44
          ^^^^^^^^^
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    45
          the equations of the binding functions
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    46
          (Trueprop equations)
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
*}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    49
text {*****************************************************}
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    50
ML {* 
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    51
(* nominal datatype parser *)
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    52
local
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    53
  structure P = OuterParse
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    54
in
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    55
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    56
val _ = OuterKeyword.keyword "bind"
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    57
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
    58
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
(* binding specification *)
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
(* should use and_list *)
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
val bind_parser = 
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    62
  P.enum "," ((P.$$$ "bind" |-- P.term) -- (P.$$$ "in" |-- P.name))
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
1223
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
    64
val constr_parser =
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
    65
  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: 1194
diff changeset
    66
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
(* datatype parser *)
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
val dt_parser =
1223
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
    69
  ((P.type_args -- P.binding -- P.opt_infix) >> P.triple1) -- 
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
    70
    (P.$$$ "=" |-- P.enum1 "|" ((constr_parser -- bind_parser -- P.opt_mixfix) >> P.triple_swap))
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    71
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    72
(* function equation parser *)
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
val fun_parser = 
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    74
  Scan.optional (P.$$$ "binder" |-- P.fixes -- SpecParse.where_alt_specs) ([],[])
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    75
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    76
(* main parser *)
1223
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
    77
val main_parser =
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: 1223
diff changeset
    78
  (P.and_list1 dt_parser) -- fun_parser
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    79
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    80
end
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
*}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
1223
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
    83
(* 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: 1194
diff changeset
    84
ML {*
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
    85
fun add_raw s = s ^ "_raw"
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
    86
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: 1194
diff changeset
    87
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
    88
fun raw_str [] s = s
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
    89
  | raw_str (s'::ss) s = 
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
    90
      if (Long_Name.base_name s) = s' 
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
    91
      then add_raw s
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
    92
      else raw_str ss s
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
    93
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
    94
fun raw_typ ty_strs (Type (a, Ts)) = Type (raw_str ty_strs a, map (raw_typ ty_strs) Ts)
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
    95
  | raw_typ ty_strs T = T  
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
    96
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
    97
fun raw_aterm trm_strs (Const (a, T)) = Const (raw_str trm_strs a, T)
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
    98
  | raw_aterm trm_strs (Free (a, T)) = Free (raw_str trm_strs a, T)
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
    99
  | raw_aterm trm_strs trm = trm
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   100
 
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   101
fun raw_term trm_strs ty_strs trm =
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   102
  trm |> Term.map_aterms (raw_aterm trm_strs) |> map_types (raw_typ ty_strs) 
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   103
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   104
fun raw_dts ty_strs dts =
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   105
let
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   106
  fun raw_dts_aux1 (bind, tys, mx) =
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   107
    (raw_bind bind, map (raw_typ ty_strs) tys, mx)
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   108
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   109
  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: 1194
diff changeset
   110
    (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: 1194
diff changeset
   111
in
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   112
  map raw_dts_aux2 dts
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   113
end
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   114
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   115
fun get_constr_strs dts =
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   116
  flat (map (fn (_, _, _, constrs) => map (fn (bn, _, _) => Binding.name_of bn) constrs) dts)
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   117
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   118
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: 1194
diff changeset
   119
  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: 1194
diff changeset
   120
*}
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   121
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   122
ML {*
1228
c179ad9d2446 declarartion of the raw datatype already works; raw binding functions throw an exception about mutual recursive types
Christian Urban <urbanc@in.tum.de>
parents: 1223
diff changeset
   123
fun raw_dts_decl dt_names dts lthy =
1223
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   124
let
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   125
  val conf = Datatype.default_config
1228
c179ad9d2446 declarartion of the raw datatype already works; raw binding functions throw an exception about mutual recursive types
Christian Urban <urbanc@in.tum.de>
parents: 1223
diff changeset
   126
1223
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   127
  val dt_names' = map add_raw dt_names
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   128
  val dts' = raw_dts dt_names dts
1228
c179ad9d2446 declarartion of the raw datatype already works; raw binding functions throw an exception about mutual recursive types
Christian Urban <urbanc@in.tum.de>
parents: 1223
diff changeset
   129
in
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: 1223
diff changeset
   130
  Local_Theory.theory_result (Datatype.add_datatype conf dt_names' dts') lthy
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: 1223
diff changeset
   131
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: 1223
diff 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: 1223
diff changeset
   133
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: 1223
diff changeset
   134
ML {*
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: 1223
diff changeset
   135
fun raw_bn_fun_decl dt_names cnstr_names bn_funs bn_eqs lthy =
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: 1223
diff changeset
   136
let
1223
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   137
  val bn_fun_strs = 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: 1194
diff changeset
   138
  
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   139
  val bn_funs' = map (fn (bn, opt_ty, mx) => 
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   140
    (raw_bind bn, Option.map (raw_typ dt_names) opt_ty, mx)) bn_funs
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   141
  val bn_eqs' = map (fn trm => 
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: 1223
diff changeset
   142
    (Attrib.empty_binding, raw_term (cnstr_names @ bn_fun_strs) dt_names trm)) bn_eqs
1223
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   143
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: 1223
diff changeset
   144
  (*Primrec.add_primrec bn_funs' bn_eqs' lthy*)
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: 1223
diff changeset
   145
  ((), lthy)
1223
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   146
end 
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   147
*}
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   148
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: 1223
diff changeset
   149
ML {* 
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: 1223
diff changeset
   150
fun nominal_datatype2 dts bn_funs bn_eqs lthy =
1223
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   151
let
1228
c179ad9d2446 declarartion of the raw datatype already works; raw binding functions throw an exception about mutual recursive types
Christian Urban <urbanc@in.tum.de>
parents: 1223
diff changeset
   152
  val dts_names = map (fn (_, s, _, _) => Binding.name_of s) dts
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: 1223
diff changeset
   153
  val ctrs_names = get_constr_strs dts
1223
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   154
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: 1223
diff changeset
   155
  lthy
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: 1223
diff changeset
   156
  |> raw_dts_decl dts_names dts 
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: 1223
diff changeset
   157
  ||>> raw_bn_fun_decl dts_names ctrs_names bn_funs bn_eqs
978
b44592adf235 Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 973
diff changeset
   158
end
b44592adf235 Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 973
diff changeset
   159
*}
b44592adf235 Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 973
diff changeset
   160
b44592adf235 Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 973
diff changeset
   161
ML {*
1228
c179ad9d2446 declarartion of the raw datatype already works; raw binding functions throw an exception about mutual recursive types
Christian Urban <urbanc@in.tum.de>
parents: 1223
diff changeset
   162
fun nominal_datatype2_cmd (dt_strs, (bn_fun_strs, bn_eq_strs)) lthy =
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   163
let
1228
c179ad9d2446 declarartion of the raw datatype already works; raw binding functions throw an exception about mutual recursive types
Christian Urban <urbanc@in.tum.de>
parents: 1223
diff changeset
   164
  val lthy_tmp = lthy
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: 1223
diff changeset
   165
    |> (Local_Theory.theory
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: 1223
diff changeset
   166
        (Sign.add_types (map (fn ((tvs, tname, mx), _) =>
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: 1223
diff changeset
   167
          (tname, length tvs, mx)) dt_strs)))
978
b44592adf235 Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 973
diff changeset
   168
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: 1223
diff changeset
   169
  fun prep_cnstr lthy (((cname, atys), mx), binders) =
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: 1223
diff changeset
   170
    (cname, map (Syntax.read_typ lthy o snd) atys, mx)
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   171
  
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: 1223
diff changeset
   172
  fun prep_dt lthy ((tvs, tname, mx), cnstrs) = 
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: 1223
diff changeset
   173
    (tvs, tname, mx, map (prep_cnstr lthy) cnstrs)
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: 1223
diff changeset
   174
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: 1223
diff changeset
   175
  fun prep_bn_fun lthy (bn, ty_str, mx) =
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: 1223
diff changeset
   176
    (bn, Option.map (Syntax.read_typ lthy) ty_str, mx) 
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   177
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: 1223
diff changeset
   178
  fun prep_bn_eq lthy (attr, trm_str) = Syntax.read_term lthy trm_str
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: 1223
diff changeset
   179
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: 1223
diff changeset
   180
  val dts_prep = map (prep_dt lthy_tmp) dt_strs
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: 1223
diff changeset
   181
  val bn_fun_prep = map (prep_bn_fun lthy_tmp) bn_fun_strs
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: 1223
diff changeset
   182
  val bn_eq_prep = map (prep_bn_eq lthy_tmp) bn_eq_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: 968
diff changeset
   183
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: 1223
diff changeset
   184
  nominal_datatype2 dts_prep bn_fun_prep [] lthy |> snd
969
cc5d18c78446 Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 968
diff changeset
   185
end
966
8ba35ce16f7e Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 962
diff changeset
   186
*}
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   187
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   188
(* Command Keyword *)
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
ML {*
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   190
let
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   191
   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
   192
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: 1223
diff changeset
   193
   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: 1223
diff changeset
   194
     (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
   195
end
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   196
*}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   197
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   198
text {* example 1 *}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   199
nominal_datatype lam =
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   200
  VAR "name"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
| APP "lam" "lam"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   202
| LET bp::"bp" t::"lam"   bind "bi bp" in t ("Let _ in _" [100,100] 100)
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   203
and bp = 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   204
  BP "name" "lam"  ("_ ::= _" [100,100] 100)
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   205
binder
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
  bi::"bp \<Rightarrow> name set"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   207
where
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   208
  "bi (BP x t) = {x}"
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: 1223
diff changeset
   209
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: 1223
diff changeset
   210
typ lam_raw
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: 1223
diff changeset
   211
term VAR_raw
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: 1223
diff changeset
   212
term BP_raw
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: 1223
diff changeset
   213
961
0f88e04eb486 Variable takes a 'name'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 954
diff changeset
   214
print_theorems
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   215
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   216
text {* examples 2 *}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   217
nominal_datatype trm =
961
0f88e04eb486 Variable takes a 'name'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 954
diff changeset
   218
  Var "name"
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   219
| App "trm" "trm"
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   220
| Lam x::"name" t::"trm"        bind x in t 
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   221
| Let p::"pat" "trm" t::"trm"   bind "f p" in t
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   222
and pat =
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   223
  PN
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   224
| PS "name"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   225
| PD "name \<times> name"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   226
binder
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   227
  f::"pat \<Rightarrow> name set"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   228
where 
978
b44592adf235 Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 973
diff changeset
   229
  "f PN = {}"
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   230
| "f (PS x) = {x}"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   231
| "f (PD (x,y)) = {x,y}"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   232
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   233
nominal_datatype trm2 =
962
8a16d6c45720 Another string in the specification.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 961
diff changeset
   234
  Var2 "name"
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   235
| App2 "trm2" "trm2"
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   236
| Lam2 x::"name" t::"trm2"          bind x in t 
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   237
| Let2 p::"pat2" "trm2" t::"trm2"   bind "f2 p" in t
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   238
and pat2 =
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   239
  PN2
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   240
| PS2 "name"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   241
| PD2 "pat2 \<times> pat2"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   242
binder
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   243
  f2::"pat2 \<Rightarrow> name set"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
where 
978
b44592adf235 Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 973
diff changeset
   245
  "f2 PN2 = {}"
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   246
| "f2 (PS2 x) = {x}"
978
b44592adf235 Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 973
diff changeset
   247
| "f2 (PD2 (p1, p2)) = (f2 p1) \<union> (f2 p2)"
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   248
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   249
text {* example type schemes *}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   250
datatype ty = 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   251
  Var "name" 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   252
| Fun "ty" "ty"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   253
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   254
nominal_datatype tyS = 
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   255
  All xs::"name list" ty::"ty" bind xs in ty
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   256
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   257
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   258
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   259
961
0f88e04eb486 Variable takes a 'name'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 954
diff changeset
   260
end
1223
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   261
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   262
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: 1223
diff changeset
   263
(* printing stuff *)
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: 1223
diff changeset
   264
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: 1223
diff changeset
   265
ML {*
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: 1223
diff changeset
   266
fun print_str_option NONE = "NONE"
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: 1223
diff changeset
   267
  | print_str_option (SOME s) = (s:bstring)
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: 1223
diff changeset
   268
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: 1223
diff changeset
   269
fun print_anno_typ lthy (NONE, ty) = Syntax.string_of_typ lthy ty
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: 1223
diff changeset
   270
  | print_anno_typ lthy (SOME x, ty) = x ^ ":" ^ Syntax.string_of_typ lthy ty
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: 1223
diff changeset
   271
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: 1223
diff changeset
   272
fun print_bind (s1, s2) = "bind " ^ s1 ^ " in " ^ s2
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: 1223
diff changeset
   273
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: 1223
diff changeset
   274
fun print_constr lthy (((name, anno_tys), bds), syn) =
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: 1223
diff changeset
   275
  (Binding.name_of name ^ " of " ^ (commas (map (print_anno_typ lthy) anno_tys)) ^ "   "
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: 1223
diff changeset
   276
   ^ (commas (map print_bind bds)) ^ "  "
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: 1223
diff changeset
   277
   ^ (Pretty.str_of (Syntax.pretty_mixfix syn)))
1223
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   278
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: 1223
diff changeset
   279
fun print_single_dt lthy (((s2, s3), syn), constrs) =
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: 1223
diff changeset
   280
   ["constructor declaration"]
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: 1223
diff changeset
   281
 @ ["type arguments: "] @ s2 
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: 1223
diff changeset
   282
 @ ["datatype name: ", Binding.name_of s3] 
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: 1223
diff changeset
   283
 @ ["typ mixfix: ", Pretty.string_of (Syntax.pretty_mixfix syn)]
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: 1223
diff changeset
   284
 @ ["constructors: "] @ (map (print_constr lthy) constrs)
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: 1223
diff changeset
   285
 |> separate "\n"
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: 1223
diff changeset
   286
 |> implode
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: 1223
diff changeset
   287
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: 1223
diff changeset
   288
fun print_single_fun_eq lthy (at, s) = 
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: 1223
diff changeset
   289
  ["function equations: ", (Syntax.string_of_term lthy s)] 
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: 1223
diff changeset
   290
  |> separate "\n"
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: 1223
diff changeset
   291
  |> implode
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: 1223
diff changeset
   292
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: 1223
diff changeset
   293
fun print_single_fun_fix lthy (b, _, _) = 
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: 1223
diff changeset
   294
  ["functions: ", Binding.name_of b] 
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: 1223
diff changeset
   295
  |> separate "\n"
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: 1223
diff changeset
   296
  |> implode
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: 1223
diff changeset
   297
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: 1223
diff changeset
   298
fun dt_name (((s2, s3), syn), constrs) = Binding.name_of s3
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: 1223
diff changeset
   299
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: 1223
diff changeset
   300
fun print_dts (dts, (funs, feqs)) lthy =
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: 1223
diff changeset
   301
let
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: 1223
diff changeset
   302
  val _ = warning (implode (separate "\n" (map (print_single_dt lthy) dts)));
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: 1223
diff changeset
   303
  val _ = warning (implode (separate "\n" (map (print_single_fun_fix lthy) funs)));
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: 1223
diff changeset
   304
  val _ = warning (implode (separate "\n" (map (print_single_fun_eq lthy) feqs)));
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: 1223
diff changeset
   305
in
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: 1223
diff changeset
   306
  ()
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: 1223
diff changeset
   307
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: 1223
diff changeset
   308
*}
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: 1223
diff changeset
   309