diff -r ac69ed8303cc -r 304bd7400a47 Nominal/ExCoreHaskell.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Nominal/ExCoreHaskell.thy Tue Mar 23 10:26:46 2010 +0100 @@ -0,0 +1,73 @@ +theory ExCoreHaskell +imports "Parser" +begin + +(* core haskell *) + +ML {* val _ = recursive := false *} + +atom_decl var +atom_decl tvar + +(* there are types, coercion types and regular types list-data-structure *) +nominal_datatype tkind = + KStar +| KFun "tkind" "tkind" +and ckind = + CKEq "ty" "ty" +and ty = + TVar "tvar" +| TC "string" +| TApp "ty" "ty" +| TFun "string" "ty_lst" +| TAll tv::"tvar" "tkind" T::"ty" bind tv in T +| TEq "ty" "ty" "ty" +and ty_lst = + TsNil +| TsCons "ty" "ty_lst" +and co = + CC "string" +| CApp "co" "co" +| CFun "string" "co_lst" +| CAll tv::"tvar" "ckind" C::"co" bind tv in C +| CEq "co" "co" "co" +| CSym "co" +| CCir "co" "co" +| CLeft "co" +| CRight "co" +| CSim "co" +| CRightc "co" +| CLeftc "co" +| CCoe "co" "co" +and co_lst = + CsNil +| CsCons "co" "co_lst" + +(* +abbreviation + "atoms A \ atom ` A" + +nominal_datatype trm = + Var "var" +| C "string" +| 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 \ kind) list" "(var \ ty) list" +binder + bv :: "pat \ atom set" +where + "bv (K s ts vs) = (atoms (set (map fst ts))) \ (atoms (set (map fst vs)))" +*) + +end + + +