Quot/Nominal/Test.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Wed, 27 Jan 2010 16:07:49 +0100
changeset 968 362402adfb88
parent 966 8ba35ce16f7e
child 969 cc5d18c78446
permissions -rw-r--r--
Undid the parsing, as it is not possible with thy->lthy interaction.
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
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
(* test for how to add datatypes *)
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
setup {*
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
Datatype.datatype_cmd
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
  ["foo"]
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
  [([], @{binding "foo"}, NoSyn, 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
    [(@{binding "zero"}, [], NoSyn),(@{binding "suc"}, ["foo"], NoSyn)])
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
  ]
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
*}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
thm foo.recs
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
thm foo.induct
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
ML{*
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
fun filtered_input str = 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
  filter OuterLex.is_proper (OuterSyntax.scan Position.none str) 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
fun parse p input = Scan.finite OuterLex.stopper (Scan.error p) input
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
*}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
(* type plus possible annotation *)
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
ML {* 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
val anno_typ =
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
  (Scan.option (OuterParse.name --| OuterParse.$$$ "::")) -- OuterParse.typ
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
*}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
ML {*
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
parse (Scan.repeat anno_typ) (filtered_input "x::string bool")
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
*}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
ML {* OuterParse.enum "," ; OuterParse.and_list1 *}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
(* binding specification *)
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
ML {* 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
datatype bind = B of string * string | BS of string * string
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
val _ = OuterKeyword.keyword "bind" 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
val _ = OuterKeyword.keyword "bindset" 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
val bind_parse_aux =
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
    (OuterParse.$$$ "bind" >> (K B)) 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
 || (OuterParse.$$$ "bindset" >> (K BS))
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
(* 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
    52
val bind_parser = 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
  OuterParse.enum ","
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
   ((bind_parse_aux -- OuterParse.term) --
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
    (OuterParse.$$$ "in" |-- OuterParse.name) >> (fn ((bind, s1), s2) => bind (s1,s2)))
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
*}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
(* datatype parser *)
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
ML {*
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
val dt_parser =
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
  OuterParse.type_args -- OuterParse.binding -- OuterParse.opt_infix -- 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
    (OuterParse.$$$ "=" |-- OuterParse.enum1 "|" 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
       (OuterParse.binding -- (Scan.repeat anno_typ) -- bind_parser -- OuterParse.opt_mixfix))
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
*}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
(* function equation parser *)
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
ML {*
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
val fun_parser = 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    69
  Scan.optional 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
    ((OuterParse.$$$ "binder") |-- OuterParse.fixes -- SpecParse.where_alt_specs) ([],[])
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
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
(* main parser *)
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
ML {*
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
val parser =
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    76
  OuterParse.and_list1 dt_parser -- fun_parser
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    77
*}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79
(* printing stuff *)
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80
ML {*
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
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
    82
  | 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
    83
*}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    84
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
ML {* 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
fun print_anno_typ (NONE, ty) = ty
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
  | print_anno_typ (SOME x, ty) = x ^ ":" ^ ty     
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
*}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
ML {*
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
fun print_bind (B (s1, s2)) = "bind " ^ s1 ^ " in " ^ s2
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
  | print_bind (BS (s1, s2)) = "bindset " ^ s1 ^ " in " ^ s2
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
*}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    95
ML {*
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    96
fun print_constr (((name, anno_tys), bds), syn) =
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
  (Binding.name_of name ^ " of " ^ (commas (map print_anno_typ anno_tys)) ^ "   " 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
   ^ (commas (map print_bind bds)) ^ "  "  
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
   ^ (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
   100
*}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   101
966
8ba35ce16f7e Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 962
diff changeset
   102
ML Local_Theory.exit_global
8ba35ce16f7e Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 962
diff changeset
   103
8ba35ce16f7e Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 962
diff changeset
   104
ML {*
8ba35ce16f7e Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 962
diff changeset
   105
fun fv_constr prefix (((name, anno_tys), bds), syn) =
8ba35ce16f7e Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 962
diff changeset
   106
  prefix ^ " (" ^ (Binding.name_of name) ^ ") = "
8ba35ce16f7e Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 962
diff changeset
   107
*}
8ba35ce16f7e Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 962
diff changeset
   108
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   109
ML {*
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
fun single_dt (((s2, s3), syn), constrs) =
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
   ["constructor declaration"]
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
 @ ["type arguments: "] @ s2 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
 @ ["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
   114
 @ ["typ mixfix: ", Pretty.string_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
   115
 @ ["constructors: "] @ (map print_constr constrs)
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
 |> separate "\n"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
 |> implode
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
*}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
966
8ba35ce16f7e Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 962
diff changeset
   120
ML {*
8ba35ce16f7e Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 962
diff changeset
   121
fun fv_dt (((s2, s3), syn), constrs) =
8ba35ce16f7e Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 962
diff changeset
   122
    map (fv_constr ("fv_" ^ Binding.name_of s3)) constrs
8ba35ce16f7e Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 962
diff changeset
   123
 |> separate "\n"
8ba35ce16f7e Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 962
diff changeset
   124
 |> implode
8ba35ce16f7e Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 962
diff changeset
   125
*}
8ba35ce16f7e Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 962
diff changeset
   126
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
ML {* fun single_fun_eq (at, s) = 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
  ["function equations: ", s] 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
  |> separate "\n"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
  |> implode
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
*}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
ML {* fun single_fun_fix (b, _, _) = 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   134
  ["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
   135
  |> separate "\n"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   136
  |> implode
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   137
*}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   138
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   139
ML {*
966
8ba35ce16f7e Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 962
diff changeset
   140
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
   141
let
968
362402adfb88 Undid the parsing, as it is not possible with thy->lthy interaction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 966
diff changeset
   142
  val _ = warning (implode (separate "\n" (map single_dt dts)))
362402adfb88 Undid the parsing, as it is not possible with thy->lthy interaction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 966
diff changeset
   143
  val _ = warning (implode (separate "\n" (map single_fun_fix funs)))
362402adfb88 Undid the parsing, as it is not possible with thy->lthy interaction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 966
diff changeset
   144
  val _ = warning (implode (separate "\n" (map single_fun_eq feqs)))
362402adfb88 Undid the parsing, as it is not possible with thy->lthy interaction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 966
diff changeset
   145
  val _ = warning (implode (separate "\n" (map fv_dt dts)))
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   146
in
968
362402adfb88 Undid the parsing, as it is not possible with thy->lthy interaction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 966
diff changeset
   147
  lthy
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   148
end   
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
*}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   150
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   151
ML {*
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   152
parser;
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   153
Datatype.datatype_cmd; 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
*}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   156
ML {*
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
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
   158
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
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
   160
  (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
   161
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   162
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
   163
 (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
   164
*}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   165
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   166
ML {*
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   167
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
   168
let
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   169
  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
   170
  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
   171
in
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
  thy'
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   173
end
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
*}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
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
   177
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
   178
  | fn_defn funs feqs lthy =
362402adfb88 Undid the parsing, as it is not possible with thy->lthy interaction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 966
diff changeset
   179
      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
   180
*}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   181
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   183
ML {*
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
fun nominal_datatype2_cmd (dts, (funs, feqs)) thy =
966
8ba35ce16f7e Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 962
diff changeset
   185
   dt_defn (map transp_dt dts) thy
968
362402adfb88 Undid the parsing, as it is not possible with thy->lthy interaction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 966
diff changeset
   186
|> Theory_Target.init NONE
966
8ba35ce16f7e Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 962
diff changeset
   187
|> fn_defn funs feqs
8ba35ce16f7e Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 962
diff changeset
   188
|> print_dts (dts, (funs, feqs))
8ba35ce16f7e Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 962
diff changeset
   189
|> Local_Theory.exit_global
8ba35ce16f7e Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 962
diff changeset
   190
*}
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   191
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   192
(* Command Keyword *)
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   193
ML {*
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   194
let
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
   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
   196
in
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   197
   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
   198
     (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
   199
end
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   200
*}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   202
text {* example 1 *}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   203
nominal_datatype lam =
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   204
  VAR "name"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   205
| APP "lam" "lam"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
| 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
   207
and bp = 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   208
  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
   209
binder
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   210
  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
   211
where
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   212
  "bi (BP x t) = {x}"
961
0f88e04eb486 Variable takes a 'name'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 954
diff changeset
   213
print_theorems
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   214
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   215
text {* examples 2 *}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   216
nominal_datatype trm =
961
0f88e04eb486 Variable takes a 'name'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 954
diff changeset
   217
  Var "name"
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   218
| App "trm" "trm"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   219
| Lam x::"name" t::"trm"        bindset x in t 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   220
| 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
   221
and pat =
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   222
  PN
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   223
| PS "name"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   224
| 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
   225
binder
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   226
  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
   227
where 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   228
  "f PN = {}"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   229
| "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
   230
| "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
   231
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   232
nominal_datatype trm2 =
962
8a16d6c45720 Another string in the specification.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 961
diff changeset
   233
  Var2 "name"
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   234
| App2 "trm2" "trm2"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   235
| Lam2 x::"name" t::"trm2"          bindset x in t 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   236
| 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
   237
and pat2 =
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   238
  PN2
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   239
| PS2 "name"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   240
| 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
   241
binder
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   242
  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
   243
where 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
  "f2 PN2 = {}"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   245
| "f2 (PS2 x) = {x}"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   246
| "f2 (PD2 (p1, p2)) = (f2 p1) \<union> (f2 p2)"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   247
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   248
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
   249
datatype ty = 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   250
  Var "name" 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   251
| Fun "ty" "ty"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   252
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   253
nominal_datatype tyS = 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   254
  All xs::"name list" ty::"ty" bindset xs in ty
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   255
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
961
0f88e04eb486 Variable takes a 'name'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 954
diff changeset
   259
end