Nominal/Ex/Term8.thy
author Christian Urban <urbanc@in.tum.de>
Fri, 27 Aug 2010 02:03:52 +0800
changeset 2438 abafea9b39bb
parent 2436 3885dc2669f9
permissions -rw-r--r--
corrected bug with fv-function generation (that was the problem with recursive binders)

theory Term8
imports "../NewParser"
begin

(* example 8 *)

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}"

(*
thm foo_bar.supp
*)

end