diff -r f78c820f67c3 -r b9e4c812d2c6 Nominal/ExCoreHaskell.thy --- a/Nominal/ExCoreHaskell.thy Sat Mar 27 14:55:07 2010 +0100 +++ b/Nominal/ExCoreHaskell.thy Sat Mar 27 16:17:45 2010 +0100 @@ -11,6 +11,7 @@ (* there are types, coercion types and regular types list-data-structure *) +ML {* val _ = alpha_type := AlphaGen *} nominal_datatype tkind = KStar | KFun "tkind" "tkind" @@ -18,18 +19,18 @@ CKEq "ty" "ty" and ty = TVar "tvar" -| TC "char" +| TC "string" | TApp "ty" "ty" -| TFun "char" "ty_lst" +| 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 "char" + CC "string" | CApp "co" "co" -| CFun "char" "co_lst" +| CFun "string" "co_lst" | CAll tv::"tvar" "ckind" C::"co" bind tv in C | CEq "co" "co" "co" | CSym "co" @@ -45,7 +46,7 @@ | CsCons "co" "co_lst" and trm = Var "var" -| C "char" +| C "string" | LAMT tv::"tvar" "tkind" t::"trm" bind tv in t | LAMC tv::"tvar" "ckind" t::"trm" bind tv in t | APP "trm" "ty" @@ -58,7 +59,7 @@ ANil | ACons p::"pat" t::"trm" "assoc_lst" bind "bv p" in t and pat = - K "char" "tvtk_lst" "tvck_lst" "vt_lst" + K "string" "tvtk_lst" "tvck_lst" "vt_lst" and vt_lst = VTNil | VTCons "var" "ty" "vt_lst" @@ -140,7 +141,18 @@ apply (simp_all add: rsp_pre) done -lemmas permute_bv[simp] = permute_bv_raw.simps[quot_lifted] +thm permute_bv_raw.simps[no_vars] + permute_bv_vt_raw.simps[quot_lifted] + permute_bv_tvck_raw.simps[quot_lifted] + permute_bv_tvtk_raw.simps[quot_lifted] + +lemma permute_bv_pre: + "permute_bv p (K c l1 l2 l3) = + K c (permute_bv_tvtk p l1) (permute_bv_tvck p l2) (permute_bv_vt p l3)" + by (lifting permute_bv_raw.simps) + +lemmas permute_bv[simp] = + permute_bv_pre permute_bv_vt_raw.simps[quot_lifted] permute_bv_tvck_raw.simps[quot_lifted] permute_bv_tvtk_raw.simps[quot_lifted] @@ -317,17 +329,17 @@ and a02: "\tkind1 tkind2 b. \\c. P1 c tkind1; \c. P1 c tkind2\ \ P1 b (KFun tkind1 tkind2)" and a03: "\ty1 ty2 b. \\c. P3 c ty1; \c. P3 c ty2\ \ P2 b (CKEq ty1 ty2)" and a04: "\tvar b. P3 b (TVar tvar)" - and a05: "\char b. P3 b (TC char)" + and a05: "\string b. P3 b (TC string)" and a06: "\ty1 ty2 b. \\c. P3 c ty1; \c. P3 c ty2\ \ P3 b (TApp ty1 ty2)" - and a07: "\char ty_lst b. \\c. P4 c ty_lst\ \ P3 b (TFun char ty_lst)" + and a07: "\string ty_lst b. \\c. P4 c ty_lst\ \ P3 b (TFun string ty_lst)" and a08: "\tvar tkind ty b. \\c. P1 c tkind; \c. P3 c ty; atom tvar \ b\ \ P3 b (TAll tvar tkind ty)" and a09: "\ty1 ty2 ty3 b. \\c. P3 c ty1; \c. P3 c ty2; \c. P3 c ty3\ \ P3 b (TEq ty1 ty2 ty3)" and a10: "\b. P4 b TsNil" and a11: "\ty ty_lst b. \\c. P3 c ty; \c. P4 c ty_lst\ \ P4 b (TsCons ty ty_lst)" - and a12: "\char b. P5 b (CC char)" + and a12: "\string b. P5 b (CC string)" and a13: "\co1 co2 b. \\c. P5 c co1; \c. P5 c co2\ \ P5 b (CApp co1 co2)" - and a14: "\char co_lst b. \\c. P6 c co_lst\ \ P5 b (CFun char co_lst)" + and a14: "\string co_lst b. \\c. P6 c co_lst\ \ P5 b (CFun string co_lst)" and a15: "\tvar ckind co b. \\c. P2 c ckind; \c. P5 c co; atom tvar \ b\ \ P5 b (CAll tvar ckind co)" and a16: "\co1 co2 co3 b. \\c. P5 c co1; \c. P5 c co2; \c. P5 c co3\ \ P5 b (CEq co1 co2 co3)" @@ -342,7 +354,7 @@ and a25: "\b. P6 b CsNil" and a26: "\co co_lst b. \\c. P5 c co; \c. P6 c co_lst\ \ P6 b (CsCons co co_lst)" and a27: "\var b. P7 b (Var var)" - and a28: "\char b. P7 b (C char)" + and a28: "\string b. P7 b (C string)" and a29: "\tvar tkind trm b. \\c. P1 c tkind; \c. P7 c trm; atom tvar \ b\ \ P7 b (LAMT tvar tkind trm)" and a30: "\tvar ckind trm b. \\c. P2 c ckind; \c. P7 c trm; atom tvar \ b\ @@ -357,8 +369,8 @@ and a37: "\b. P8 b ANil" and a38: "\pat trm assoc_lst b. \\c. P9 c pat; \c. P7 c trm; \c. P8 c assoc_lst; bv(pat) \* b\ \ P8 b (ACons pat trm assoc_lst)" - and a39: "\char tvtk_lst tvck_lst vt_lst b. \\c. P11 c tvtk_lst; \c. P12 c tvck_lst; \c. P10 c vt_lst\ - \ P9 b (K char tvtk_lst tvck_lst vt_lst)" + and a39: "\string tvtk_lst tvck_lst vt_lst b. \\c. P11 c tvtk_lst; \c. P12 c tvck_lst; \c. P10 c vt_lst\ + \ P9 b (K string tvtk_lst tvck_lst vt_lst)" and a40: "\b. P10 b VTNil" and a41: "\var ty vt_lst b. \\c. P3 c ty; \c. P10 c vt_lst\ \ P10 b (VTCons var ty vt_lst)" and a42: "\b. P11 b TVTKNil"