Nominal/Ex/Term8.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Mon, 10 May 2010 15:54:16 +0200
changeset 2094 56b849d348ae
parent 2082 0854af516f14
child 2097 60ce41a11180
permissions -rw-r--r--
Recursive examples with relation composition

theory Term8
imports "../NewParser"
begin

(* example 8 *)

atom_decl name

(* this should work but gives an error at the moment:
   seems like in the symmetry proof
*)

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 s in b
binder
  bv
where
  "bv (Bar0 x) = {}"
| "bv (Bar1 v x b) = {atom v}"


end