Nominal/Test.thy
author Christian Urban <urbanc@in.tum.de>
Sat, 27 Feb 2010 11:54:59 +0100
changeset 1284 212f3ab40cc2
parent 1283 6a133abb7633
child 1285 e3ac56d3b7af
permissions -rw-r--r--
streamlined parser

theory Test
imports "Parser"
begin

text {* example 1 *}

nominal_datatype lam =
  VAR "name"
| APP "lam" "lam"
| LET bp::"bp" t::"lam"   bind "bi bp" in t ("Let _ in _" [100,100] 100)
and bp = 
  BP "name" "lam"  ("_ ::= _" [100,100] 100)
binder
  bi::"bp \<Rightarrow> name set"
where
  "bi (BP x t) = {x}"

typ lam_raw
term VAR_raw
term APP_raw
term LET_raw
term Test.BP_raw
thm bi_raw.simps

print_theorems

text {* example 2 *}

nominal_datatype trm' =
  Var "name"
| App "trm'" "trm'"
| Lam x::"name" t::"trm'"          bind x in t 
| Let p::"pat'" "trm'" t::"trm'"   bind "f p" in t
and pat' =
  PN
| PS "name"
| PD "name" "name"
binder
  f::"pat' \<Rightarrow> name set"
where 
  "f PN = {}"
| "f (PS x) = {x}"
| "f (PD x y) = {x,y}"

thm f_raw.simps

nominal_datatype trm0 =
  Var0 "name"
| App0 "trm0" "trm0"
| Lam0 x::"name" t::"trm0"          bind x in t 
| Let0 p::"pat0" "trm0" t::"trm0"   bind "f0 p" in t
and pat0 =
  PN0
| PS0 "name"
| PD0 "pat0" "pat0"
binder
  f0::"pat0 \<Rightarrow> name set"
where 
  "f0 PN0 = {}"
| "f0 (PS0 x) = {x}"
| "f0 (PD0 p1 p2) = (f0 p1) \<union> (f0 p2)"

thm f0_raw.simps

text {* example type schemes *}
datatype t = 
  Var "name" 
| Fun "t" "t"

nominal_datatype tyS = 
  All xs::"name list" ty::"t" bind xs in ty



(* example 1 from Terms.thy *)

nominal_datatype trm1 =
  Vr1 "name"
| Ap1 "trm1" "trm1"
| Lm1 x::"name" t::"trm1"      bind x in t 
| Lt1 p::"bp1" "trm1" t::"trm1" bind "bv1 p" in t 
and bp1 =
  BUnit1
| BV1 "name"
| BP1 "bp1" "bp1"
binder
  bv1
where
  "bv1 (BUnit1) = {}"
| "bv1 (BV1 x) = {atom x}"
| "bv1 (BP1 bp1 bp2) = (bv1 bp1) \<union> (bv1 bp2)"

thm bv1_raw.simps

(* example 2 from Terms.thy *)

nominal_datatype trm2 =
  Vr2 "name"
| Ap2 "trm2" "trm2"
| Lm2 x::"name" t::"trm2"       bind x in t
| Lt2 r::"rassign" t::"trm2"    bind "bv2 r" in t
and rassign = 
  As "name" "trm2"
binder
  bv2
where
  "bv2 (As x t) = {atom x}"

(* example 3 from Terms.thy *)

nominal_datatype trm3 =
  Vr3 "name"
| Ap3 "trm3" "trm3"
| Lm3 x::"name" t::"trm3"        bind x in t
| Lt3 r::"rassigns3" t::"trm3"   bind "bv3 r" in t
and rassigns3 =
  ANil
| ACons "name" "trm3" "rassigns3"
binder
  bv3
where
  "bv3 ANil = {}"
| "bv3 (ACons x t as) = {atom x} \<union> (bv3 as)"

(* example 4 from Terms.thy *)

nominal_datatype trm4 =
  Vr4 "name"
| Ap4 "trm4" "trm4 list"
| Lm4 x::"name" t::"trm4"  bind x in t

(* example 5 from Terms.thy *)

nominal_datatype trm5 =
  Vr5 "name"
| Ap5 "trm5" "trm5"
| Lt5 l::"lts" t::"trm5"  bind "bv5 l" in t
and lts =
  Lnil
| Lcons "name" "trm5" "lts"
binder
  bv5
where
  "bv5 Lnil = {}"
| "bv5 (Lcons n t ltl) = {atom n} \<union> (bv5 ltl)"

(* example 6 from Terms.thy *)

nominal_datatype trm6 =
  Vr6 "name"
| Lm6 x::"name" t::"trm6"         bind x in t
| Lt6 left::"trm6" right::"trm6"  bind "bv6 left" in right
binder
  bv6
where
  "bv6 (Vr6 n) = {}"
| "bv6 (Lm6 n t) = {atom n} \<union> bv6 t"
| "bv6 (Lt6 l r) = bv6 l \<union> bv6 r"

(* example 7 from Terms.thy *)

nominal_datatype trm7 =
  Vr7 "name"
| Lm7 l::"name" r::"trm7"   bind l in r
| Lt7 l::"trm7" r::"trm7"   bind "bv7 l" in r
binder 
  bv7 
where
  "bv7 (Vr7 n) = {atom n}"
| "bv7 (Lm7 n t) = bv7 t - {atom n}"
| "bv7 (Lt7 l r) = bv7 l \<union> bv7 r"

(* example 8 from Terms.thy *)

nominal_datatype foo8 =
  Foo0 "name"
| Foo1 b::"bar8" f::"foo8" bind "bv8 b" in f --"check fo error if this is called foo"
and bar8 =
  Bar0 "name"
| Bar1 "name" s::"name" b::"bar8" bind s in b
binder 
  bv8
where
  "bv8 (Bar0 x) = {}"
| "bv8 (Bar1 v x b) = {atom v}"

(* example 9 from Terms.thy *)

nominal_datatype lam9 =
  Var9 "name"
| Lam9 n::"name" l::"lam9" bind n in l
and bla9 =
  Bla9 f::"lam9" s::"lam9" bind "bv9 f" in s
binder
  bv9
where
  "bv9 (Var9 x) = {}"
| "bv9 (Lam9 x b) = {atom x}"

(* example form Leroy 96 about modules *)


nominal_datatype mexp =
  Acc path
| Stru body
| Funct x::name sexp m::mexp    bind x in m
| FApp mexp path
| Ascr mexp sexp
and body =
  Empty
| Seq c::defn d::body bind "cbinders c" in d
and defn =  
  Type name tyty
| Dty name
| DStru name mexp
| Val name trmtrm
and sexp =
  Sig sbody
| SFunc name sexp sexp
and sbody = 
  SEmpty
| SSeq C::spec D::sbody  bind "Cbinders C" in D
and spec =
  Type1 name 
| Type2 name tyty
| SStru name sexp
| SVal name tyty
and tyty =
  Tyref1 name
| Tyref2 path tyty
| Fun tyty tyty
and path =
  Sref1 name
| Sref2 path name
and trmtrm =
  Tref1 name
| Tref2 path name
| Lam v::name tyty M::trmtrm  bind v in M
| App trmtrm trmtrm
| Let body trmtrm
binder
    cbinders :: "defn \<Rightarrow> atom set"
and Cbinders :: "spec \<Rightarrow> atom set"
where
  "cbinders (Type t T) = {atom t}"
| "cbinders (Dty t) = {atom t}"
| "cbinders (DStru x s) = {atom x}"
| "cbinders (Val v M) = {atom v}"
| "Cbinders (Type1 t) = {atom t}"
| "Cbinders (Type2 t T) = {atom t}"
| "Cbinders (SStru x S) = {atom x}"
| "Cbinders (SVal v T) = {atom v}"  

(* core haskell *)

atom_decl var
atom_decl tvar
atom_decl co

datatype sort = 
  TY tvar
| CO co

datatype kind = 
  KStar
| KFun kind kind
| KEq kind kind

(* there are types, coercion types and regular types *)
nominal_datatype ty =
  TVar tvar
| TFun string "ty list"
| TAll tvar kind ty --"some binding"
| TSym ty
| TCir ty ty
| TApp ty ty
| TLeft ty
| TRight ty
| TEq ty
| TRightc ty
| TLeftc ty
| TCoe ty ty

typedecl ty --"hack since ty is not yet defined"

abbreviation 
  "atoms A \<equiv> atom ` A"

nominal_datatype trm =
  Var var
| LAM tv::tvar kind t::trm   bind tv in t 
| APP trm ty
| Lam v::var ty t::trm       bind v in t
| App trm trm
| Let x::var ty trm t::trm   bind x in t
| Case trm "assoc list"
| Cast trm ty   --"ty is supposed to be a coercion type only"
and assoc = 
  A p::pat t::trm bind "bv p" in t 
and pat = 
  K string "(tvar \<times> kind) list" "(var \<times> ty) list"
binder
 bv :: "pat \<Rightarrow> atom set"
where
 "bv (K s ts vs) = (atoms (set (map fst ts))) \<union> (atoms (set (map fst vs)))"


thm bv_raw.simps

end