Nominal/Ex/Term8.thy
author Christian Urban <urbanc@in.tum.de>
Tue, 04 May 2010 14:25:22 +0100
changeset 2046 73c50e913db6
parent 2045 6800fcaafa2a
child 2082 0854af516f14
permissions -rw-r--r--
tuned and added some comments to the code; added also an exception for early exit of the nominal2_cmd function

theory Term8
imports "../NewParser"
begin

(* example 8 from Terms.thy *)

atom_decl name

nominal_datatype foo =
  Foo0 "name"
| Foo1 b::"bar" f::"foo" bind_set "bv b" in f
and bar =
  Bar0 "name"
| Bar1 "name" s::"name" b::"bar" bind_set s in b
binder
  bv
where
  "bv (Bar0 x) = {}"
| "bv (Bar1 v x b) = {atom v}"

end