Quot/Nominal/Test.thy
author Christian Urban <urbanc@in.tum.de>
Tue, 23 Feb 2010 13:32:35 +0100
changeset 1223 160343d86a6f
parent 1194 3d54fcc5f41a
child 1228 c179ad9d2446
permissions -rw-r--r--
"raw"-ified the term-constructors and types given in the specification
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 =
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    78
  P.and_list1 dt_parser -- fun_parser
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 {*
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   123
fun nominal_dt_decl dt_names dts bn_funs bn_eqs lthy =
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
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   126
  
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
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   129
  val constr_strs = 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
   130
  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
   131
  
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   132
  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
   133
    (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
   134
  val bn_eqs' = map (fn trm => 
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   135
    (Attrib.empty_binding, raw_term (constr_strs @ bn_fun_strs) dt_names trm)) bn_eqs
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   136
in
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   137
  lthy
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   138
  |> (Local_Theory.theory_result (Datatype.add_datatype conf dt_names' dts'))
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   139
  ||>> Primrec.add_primrec bn_funs' bn_eqs'
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   140
end 
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   141
*}
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   142
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   143
local_setup {*
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   144
let
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   145
  val T0 = Type ("Test.foo", [])
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   146
  val T1 = T0 --> @{typ "nat"}
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   147
in
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   148
nominal_dt_decl
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   149
  ["foo"]
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   150
  [([], @{binding "foo"}, NoSyn, 
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   151
    [(@{binding "myzero"}, [], NoSyn),(@{binding "mysuc"}, [Type ("Test.foo", [])], NoSyn)])]
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   152
  [(@{binding "Bar"}, SOME T1, NoSyn)]
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   153
  [HOLogic.mk_Trueprop (HOLogic.mk_eq 
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   154
        (Free ("Bar", T1) $ Const ("Test.foo_raw.myzero", T0), @{term "0::nat"})),
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   155
   HOLogic.mk_Trueprop (HOLogic.mk_eq 
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   156
        (Free ("Bar", T1) $ (Const ("Test.foo_raw.mysuc", T0 --> T0) $ Free ("n", T0)), @{term "0::nat"}))]
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   157
  #> snd
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   158
end
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   159
*}
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   160
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   161
typ foo_raw
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   162
thm foo_raw.induct
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   163
term myzero_raw
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   164
term mysuc_raw
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   165
thm Bar_raw.simps
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   166
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   167
(* printing stuff *)
1223
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   168
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   169
ML {*
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   170
fun print_str_option NONE = "NONE"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   171
  | print_str_option (SOME s) = (s:bstring)
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
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
   173
fun print_anno_typ lthy (NONE, ty) = Syntax.string_of_typ lthy ty
cc5d18c78446 Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 968
diff changeset
   174
  | print_anno_typ lthy (SOME x, ty) = x ^ ":" ^ Syntax.string_of_typ lthy ty
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   176
fun print_bind (s1, s2) = "bind " ^ s1 ^ " in " ^ s2
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
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
   178
fun print_constr lthy (((name, anno_tys), bds), syn) =
cc5d18c78446 Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 968
diff changeset
   179
  (Binding.name_of name ^ " of " ^ (commas (map (print_anno_typ lthy) anno_tys)) ^ "   "
cc5d18c78446 Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 968
diff changeset
   180
   ^ (commas (map print_bind bds)) ^ "  "
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   181
   ^ (Pretty.str_of (Syntax.pretty_mixfix syn)))
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   183
fun print_single_dt lthy (((s2, s3), syn), constrs) =
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
   ["constructor declaration"]
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
 @ ["type arguments: "] @ s2 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   186
 @ ["datatype name: ", Binding.name_of s3] 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   187
 @ ["typ mixfix: ", Pretty.string_of (Syntax.pretty_mixfix syn)]
969
cc5d18c78446 Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 968
diff changeset
   188
 @ ["constructors: "] @ (map (print_constr lthy) constrs)
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
 |> separate "\n"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   190
 |> implode
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   191
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   192
fun print_single_fun_eq lthy (at, s) = 
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
   193
  ["function equations: ", (Syntax.string_of_term lthy s)] 
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   194
  |> separate "\n"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
  |> implode
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   196
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   197
fun print_single_fun_fix lthy (b, _, _) = 
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   198
  ["functions: ", Binding.name_of b] 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   199
  |> separate "\n"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   200
  |> implode
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
970
c16135580a45 Some processing of variables in constructors to get free variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 969
diff changeset
   202
fun dt_name (((s2, s3), syn), constrs) = Binding.name_of s3
c16135580a45 Some processing of variables in constructors to get free variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 969
diff changeset
   203
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
   204
fun print_dts (dts, (funs, feqs)) lthy =
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   205
let
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   206
  val _ = warning (implode (separate "\n" (map (print_single_dt lthy) dts)));
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   207
  val _ = warning (implode (separate "\n" (map (print_single_fun_fix lthy) funs)));
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   208
  val _ = warning (implode (separate "\n" (map (print_single_fun_eq lthy) feqs)));
978
b44592adf235 Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 973
diff changeset
   209
in
b44592adf235 Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 973
diff changeset
   210
  ()
b44592adf235 Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 973
diff changeset
   211
end
b44592adf235 Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 973
diff changeset
   212
*}
b44592adf235 Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 973
diff changeset
   213
b44592adf235 Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 973
diff changeset
   214
ML {*
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   215
fun transp_ty_arg (anno, ty) = ty
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   216
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   217
fun transp_constr (((constr_name, ty_args), bind), mx) = 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   218
  (constr_name, map transp_ty_arg ty_args, mx)
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   219
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   220
fun transp_dt (((ty_args, ty_name), mx), constrs) = 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   221
 (ty_args, ty_name, mx, map transp_constr constrs)
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   222
*}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   223
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   224
ML {*
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   225
fun dt_defn dts thy =
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   226
let
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   227
  val names = map (fn (_, s, _, _) => Binding.name_of s) dts
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   228
  val thy' = Datatype.datatype_cmd names dts thy
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   229
in
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   230
  thy'
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   231
end
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
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   234
ML {*
968
362402adfb88 Undid the parsing, as it is not possible with thy->lthy interaction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 966
diff changeset
   235
fun fn_defn [] [] lthy = lthy
362402adfb88 Undid the parsing, as it is not possible with thy->lthy interaction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 966
diff changeset
   236
  | fn_defn funs feqs lthy =
978
b44592adf235 Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 973
diff changeset
   237
      Function_Fun.add_fun Function_Common.default_config funs feqs false lthy
b44592adf235 Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 973
diff changeset
   238
b44592adf235 Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 973
diff changeset
   239
fun fn_defn_cmd [] [] lthy = lthy
b44592adf235 Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 973
diff changeset
   240
  | fn_defn_cmd funs feqs lthy =
b44592adf235 Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 973
diff changeset
   241
      Function_Fun.add_fun_cmd Function_Common.default_config funs feqs false lthy
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   242
*}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   243
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
ML {*
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   245
fun parse_fun lthy (b, NONE, m) = (b, NONE, m)
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   246
  | parse_fun lthy (b, SOME ty, m) = (b, SOME (Syntax.read_typ lthy ty), m)
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   247
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   248
fun parse_feq lthy (b, t) = (b, Syntax.read_prop lthy t);
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   249
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   250
fun parse_anno_ty lthy (anno, ty) = (anno, Syntax.read_typ lthy ty);
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   251
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   252
fun parse_constr lthy (((constr_name, ty_args), bind), mx) =
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   253
  (((constr_name, map (parse_anno_ty lthy) ty_args), bind), mx);
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   254
  
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   255
fun parse_dt lthy (((ty_args, ty_name), mx), constrs) =
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   256
   (((ty_args, ty_name), mx), map (parse_constr lthy) constrs);
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   257
*}
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   258
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   259
ML {*
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   260
fun nominal_datatype2_cmd (dt_strs, (fun_strs, feq_strs)) thy =
969
cc5d18c78446 Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 968
diff changeset
   261
let
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   262
  val thy' = dt_defn (map transp_dt dt_strs) thy
969
cc5d18c78446 Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 968
diff changeset
   263
  val lthy = Theory_Target.init NONE thy'
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   264
  val lthy' = fn_defn_cmd fun_strs feq_strs lthy
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   265
  val funs = map (parse_fun lthy') fun_strs
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   266
  val feqs = map (parse_feq lthy') feq_strs
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   267
  val dts = map (parse_dt lthy') dt_strs
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   268
  val _ = print_dts (dts, (funs, feqs)) lthy'
969
cc5d18c78446 Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 968
diff changeset
   269
in
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   270
  Local_Theory.exit_global lthy'
969
cc5d18c78446 Parsing of the input as terms and types, and passing them as such to the function package.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 968
diff changeset
   271
end
966
8ba35ce16f7e Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 962
diff changeset
   272
*}
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   273
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   274
(* Command Keyword *)
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   275
ML {*
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   276
let
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   277
   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
   278
in
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   279
   OuterSyntax.command "nominal_datatype" "test" kind 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   280
     (parser >> (Toplevel.theory o nominal_datatype2_cmd))
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   281
end
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   282
*}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   283
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   284
text {* example 1 *}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   285
nominal_datatype lam =
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   286
  VAR "name"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   287
| APP "lam" "lam"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   288
| 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
   289
and bp = 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   290
  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
   291
binder
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   292
  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
   293
where
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   294
  "bi (BP x t) = {x}"
961
0f88e04eb486 Variable takes a 'name'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 954
diff changeset
   295
print_theorems
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   296
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   297
text {* examples 2 *}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   298
nominal_datatype trm =
961
0f88e04eb486 Variable takes a 'name'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 954
diff changeset
   299
  Var "name"
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   300
| App "trm" "trm"
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   301
| 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
   302
| 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
   303
and pat =
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   304
  PN
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   305
| PS "name"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   306
| 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
   307
binder
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   308
  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
   309
where 
978
b44592adf235 Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 973
diff changeset
   310
  "f PN = {}"
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   311
| "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
   312
| "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
   313
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   314
nominal_datatype trm2 =
962
8a16d6c45720 Another string in the specification.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 961
diff changeset
   315
  Var2 "name"
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   316
| App2 "trm2" "trm2"
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   317
| 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
   318
| 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
   319
and pat2 =
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   320
  PN2
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   321
| PS2 "name"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   322
| 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
   323
binder
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   324
  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
   325
where 
978
b44592adf235 Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 973
diff changeset
   326
  "f2 PN2 = {}"
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   327
| "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
   328
| "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
   329
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   330
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
   331
datatype ty = 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   332
  Var "name" 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   333
| Fun "ty" "ty"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   334
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   335
nominal_datatype tyS = 
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   336
  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
   337
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   338
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   339
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   340
961
0f88e04eb486 Variable takes a 'name'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 954
diff changeset
   341
end
1223
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   342
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   343
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   344
(* probably can be done with a clever morphism trick *)
160343d86a6f "raw"-ified the term-constructors and types given in the specification
Christian Urban <urbanc@in.tum.de>
parents: 1194
diff changeset
   345