Quot/Examples/Terms.thy
author Christian Urban <urbanc@in.tum.de>
Thu, 14 Jan 2010 12:14:35 +0100
changeset 872 2605ea41bbdd
child 874 ab8a58973861
permissions -rw-r--r--
added 3 calculi with interesting binding structure

theory Terms
imports Nominal "../QuotMain" "../QuotList"
begin

atom_decl name

(* lets with binding patterns *)
datatype trm1 =
  Vr1 "name"
| Ap1 "trm1" "trm1"
| Lm1 "name" "trm1"
| Lt1 "bp" "trm1" "trm1"
and bp =
  BUnit
| BVr "name"
| BPr "bp" "bp"

(* lets with single assignments *)
datatype trm2 =
  Vr2 "name"
| Ap2 "trm2" "trm2"
| Lm2 "name" "trm2"
| Lt2 "assign" "trm2"
and assign =
  As "name" "trm2"


(* lets with many assignments *)
datatype trm3 =
  Vr3 "name"
| Ap3 "trm3" "trm3"
| Lm3 "name" "trm3"
| Lt3 "assign" "trm3"
and assigns =
  ANil
| ACons "name" "trm3" "assigns"