Quot/Nominal/Test.thy
author Christian Urban <urbanc@in.tum.de>
Thu, 18 Feb 2010 23:07:28 +0100
changeset 1194 3d54fcc5f41a
parent 1087 bb7f4457091a
child 1223 160343d86a6f
permissions -rw-r--r--
start work with the parser
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
datatype simple = Foo
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
local_setup {* 
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    23
Primrec.add_primrec [(@{binding "Bar"}, SOME @{typ "simple \<Rightarrow> nat"}, NoSyn)]
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    24
 [(Attrib.empty_binding, HOLogic.mk_Trueprop @{term "Bar Foo = Suc 0"})] #> snd
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
thm Bar.simps
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
lemma "Bar Foo = XXX"
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    29
apply(simp)
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    30
oops
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
section {* test for setting up a datatype on the ML-level *}
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
setup {*
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    34
Datatype.add_datatype Datatype.default_config
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
  ["foo"]
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
  [([], @{binding "foo"}, NoSyn, 
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    37
    [(@{binding "zero"}, [], NoSyn),(@{binding "suc"}, [Type ("Test.foo", [])], NoSyn)])
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    38
  ] #> snd
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    39
*}
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    40
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    41
ML {*
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    42
  PolyML.makestring (#descr (Datatype.the_info @{theory} "Test.foo"))
954
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
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
thm foo.recs
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
thm foo.induct
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    48
(* adds "_raw" to the end of constants and types *)
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    49
ML {*
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    50
fun raw_str [] s = s
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    51
  | raw_str (s'::ss) s = (if s = s' then s ^ "_str" else raw_str ss s)
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    53
val raw_strs = map raw_str
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    54
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    55
fun raw_typ ty_strs (Type (a, Ts)) = Type (raw_str ty_strs a, map (raw_typ ty_strs) Ts)
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    56
  | raw_typ ty_strs T = T  
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    58
fun raw_term trm_strs ty_strs (Const (a, T)) = Const (raw_str trm_strs a, raw_typ ty_strs T)
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    59
  | raw_term trm_strs ty_strs (Abs (a, T, t)) = Abs (a, T, raw_term trm_strs ty_strs t)
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    60
  | raw_term trm_strs ty_strs (t $ u) = (raw_term trm_strs ty_strs t) $ (raw_term trm_strs ty_strs u)
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    61
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    62
fun raw_binding bn = Binding.suffix_name "_raw" bn
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
*}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    65
section{* Interface for nominal_datatype *}
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    66
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    67
text {*
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    68
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    69
Nominal-Datatype-part:
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    70
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    71
1st Arg: string list  
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    72
         ^^^^^^^^^^^
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    73
         strings of the types to be defined
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    74
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    75
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
    76
         ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^   ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    77
               type(s) to be defined             constructors list
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    78
               (ty args, name, syn)              (name, typs, syn)
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
Binder-Function-part:
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    81
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    82
3rd Arg: (binding * typ option * mixfix) list 
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    83
         ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^    
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    84
            binding function(s)           
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    85
              to be defined               
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    86
            (name, type, syn)             
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    87
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    88
4th Arg: (attrib * term) list 
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    89
          ^^^^^^^^^^^^^^^^^^^
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    90
          the equations of the binding functions
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    91
          (Trueprop equations)
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    92
954
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
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    95
ML {* Datatype.add_datatype *}
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    96
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
ML {*
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    98
fun raw_definitions dt_names dts bn_funs bn_eqs lthy =
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
    99
let
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   100
  val conf = Datatype.default_config
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   101
  val bn_funs' = map (fn (bn, ty) => (raw_binding bn, ty, NoSyn)) bn_funs
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   102
  val bn_eqs' = map (pair Attrib.empty_binding) bn_eqs
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   103
in 
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   104
  lthy
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   105
  |> (Local_Theory.theory_result (Datatype.add_datatype conf dt_names dts))
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   106
  ||>> Primrec.add_primrec bn_funs' bn_eqs'
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   107
end 
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
*}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   109
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   110
ML {* @{binding "zero"} *}
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   111
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   112
local_setup {*
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   113
raw_definitions
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   114
  ["foo'"]
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   115
  [([], @{binding "foo'"}, NoSyn, 
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   116
    [(@{binding "zero'"}, [], NoSyn),(@{binding "suc'"}, [Type ("Test.foo'", [])], NoSyn)])]
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   117
  [(@{binding "Bar'"}, SOME @{typ "simple \<Rightarrow> nat"})]
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   118
  [HOLogic.mk_Trueprop @{term "Bar'_raw Foo = Suc 0"}]
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   119
  #> snd
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   120
*}
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   121
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   122
typ foo'
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   123
thm Bar'.simps
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   124
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   125
text {*****************************************************}
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   126
ML {* 
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   127
(* nominal datatype parser *)
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   128
local
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   129
  structure P = OuterParse
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   130
in
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   131
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   132
val _ = OuterKeyword.keyword "bind"
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   133
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
   134
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   135
(* binding specification *)
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   136
(* 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
   137
val bind_parser = 
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   138
  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
   139
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   140
(* datatype parser *)
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   141
val dt_parser =
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   142
  P.type_args -- P.binding -- P.opt_infix -- 
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   143
    (P.$$$ "=" |-- P.enum1 "|" (P.binding -- (Scan.repeat anno_typ) -- bind_parser -- P.opt_mixfix))
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   144
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   145
(* function equation parser *)
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   146
val fun_parser = 
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   147
  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
   148
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   149
(* main parser *)
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   150
val parser =
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   151
  P.and_list1 dt_parser -- fun_parser
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   152
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   153
end
954
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
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
(* printing stuff *)
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
ML {*
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
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
   160
  | 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
   161
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
   162
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
   163
  | 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
   164
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   165
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
   166
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
   167
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
   168
  (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
   169
   ^ (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
   170
   ^ (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
   171
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   172
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
   173
   ["constructor declaration"]
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
 @ ["type arguments: "] @ s2 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
 @ ["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
   176
 @ ["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
   177
 @ ["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
   178
 |> separate "\n"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   179
 |> implode
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   180
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   181
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
   182
  ["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
   183
  |> separate "\n"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
  |> implode
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   186
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
   187
  ["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
   188
  |> separate "\n"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
  |> implode
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   190
970
c16135580a45 Some processing of variables in constructors to get free variables.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 969
diff changeset
   191
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
   192
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
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
   194
let
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   195
  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
   196
  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
   197
  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
   198
in
b44592adf235 Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 973
diff changeset
   199
  ()
b44592adf235 Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 973
diff changeset
   200
end
b44592adf235 Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 973
diff changeset
   201
*}
b44592adf235 Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 973
diff changeset
   202
b44592adf235 Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 973
diff changeset
   203
ML {*
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   204
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
   205
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
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
   207
  (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
   208
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   209
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
   210
 (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
   211
*}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   212
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   213
ML {*
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   214
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
   215
let
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   216
  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
   217
  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
   218
in
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   219
  thy'
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   220
end
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   221
*}
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
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
   224
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
   225
  | 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
   226
      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
   227
b44592adf235 Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 973
diff changeset
   228
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
   229
  | 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
   230
      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
   231
*}
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
ML {*
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   234
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
   235
  | 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
   236
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   237
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
   238
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   239
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
   240
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   241
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
   242
  (((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
   243
  
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   244
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
   245
   (((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
   246
*}
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
ML {*
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   249
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
   250
let
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   251
  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
   252
  val lthy = Theory_Target.init NONE thy'
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   253
  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
   254
  val funs = map (parse_fun lthy') fun_strs
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   255
  val feqs = map (parse_feq lthy') feq_strs
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   256
  val dts = map (parse_dt lthy') dt_strs
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   257
  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
   258
in
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   259
  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
   260
end
966
8ba35ce16f7e Some cleaning of thy vs lthy vs context.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 962
diff changeset
   261
*}
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   262
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   263
(* Command Keyword *)
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   264
ML {*
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   265
let
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   266
   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
   267
in
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   268
   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
   269
     (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
   270
end
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   271
*}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   272
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   273
text {* example 1 *}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   274
nominal_datatype lam =
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   275
  VAR "name"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   276
| APP "lam" "lam"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   277
| 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
   278
and bp = 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   279
  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
   280
binder
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   281
  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
   282
where
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   283
  "bi (BP x t) = {x}"
961
0f88e04eb486 Variable takes a 'name'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 954
diff changeset
   284
print_theorems
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   285
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   286
text {* examples 2 *}
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   287
nominal_datatype trm =
961
0f88e04eb486 Variable takes a 'name'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 954
diff changeset
   288
  Var "name"
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   289
| App "trm" "trm"
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   290
| 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
   291
| 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
   292
and pat =
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   293
  PN
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   294
| PS "name"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   295
| 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
   296
binder
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   297
  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
   298
where 
978
b44592adf235 Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 973
diff changeset
   299
  "f PN = {}"
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   300
| "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
   301
| "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
   302
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   303
nominal_datatype trm2 =
962
8a16d6c45720 Another string in the specification.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 961
diff changeset
   304
  Var2 "name"
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   305
| App2 "trm2" "trm2"
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   306
| 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
   307
| 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
   308
and pat2 =
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   309
  PN2
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   310
| PS2 "name"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   311
| 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
   312
binder
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   313
  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
   314
where 
978
b44592adf235 Improper interface for datatype and function packages and proper interface lateron.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 973
diff changeset
   315
  "f2 PN2 = {}"
954
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   316
| "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
   317
| "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
   318
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   319
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
   320
datatype ty = 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   321
  Var "name" 
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   322
| Fun "ty" "ty"
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   323
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   324
nominal_datatype tyS = 
1194
3d54fcc5f41a start work with the parser
Christian Urban <urbanc@in.tum.de>
parents: 1087
diff changeset
   325
  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
   326
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   327
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   328
c009d2535896 very rough example file for how nominal2 specification can be parsed
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   329
961
0f88e04eb486 Variable takes a 'name'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 954
diff changeset
   330
end