diff -r b5141d1ab24f -r 58abc09d6f9c Nominal/ExCoreHaskell.thy --- a/Nominal/ExCoreHaskell.thy Mon Mar 29 17:32:17 2010 +0200 +++ b/Nominal/ExCoreHaskell.thy Mon Mar 29 18:12:54 2010 +0200 @@ -323,7 +323,9 @@ apply (simp add: finb1 finb2 finb3) done -lemma strong_inudction_principle: +thm tkind_ckind_ty_ty_lst_co_co_lst_trm_assoc_lst_pat_vars_tvars_cvars.induct + +lemma strong_induction_principle: assumes a01: "\b. P1 b KStar" 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)" @@ -333,32 +335,36 @@ 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 TvsNil" - and a11: "\ty ty_lst b. \\c. P3 c ty; \c. P4 c ty_lst\ \ P4 b (TvsCons ty ty_lst)" - and a12: "\string b. P5 b (CC string)" + and a09: "\ck ty b. \\c. P2 c ck; \c. P3 c ty\ \ P3 b (TEq ck ty)" + 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: "\string b. P5 b (CVar string)" + and a12a:"\str b. P5 b (CConst str)" and a13: "\co1 co2 b. \\c. P5 c co1; \c. P5 c co2\ \ P5 b (CApp co1 co2)" 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)" - and a17: "\co b. \\c. P5 c co\ \ P5 b (CSym co)" + and a16: "\ck co b. \\c. P2 c ck; \c. P5 c co\ \ P5 b (CEq ck co)" + and a17: "\ty b. \\c. P3 c ty\ \ P5 b (CRefl ty)" + and a17a: "\co b. \\c. P5 c co\ \ P5 b (CSym co)" and a18: "\co1 co2 b. \\c. P5 c co1; \c. P5 c co2\ \ P5 b (CCir co1 co2)" + and a18a:"\co ty b. \\c. P5 c co; \c. P3 c ty\ \ P5 b (CAt co ty)" and a19: "\co b. \\c. P5 c co\ \ P5 b (CLeft co)" and a20: "\co b. \\c. P5 c co\ \ P5 b (CRight co)" - and a21: "\co b. \\c. P5 c co\ \ P5 b (CSim co)" + and a21: "\co1 co2 b. \\c. P5 c co1; \c. P5 c co2\ \ P5 b (CSim co1 co2)" and a22: "\co b. \\c. P5 c co\ \ P5 b (CRightc co)" and a23: "\co b. \\c. P5 c co\ \ P5 b (CLeftc co)" and a24: "\co1 co2 b. \\c. P5 c co1; \c. P5 c co2\ \ P5 b (CCoe co1 co2)" 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: "\string b. P7 b (C string)" + and a28: "\string b. P7 b (K 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\ \ P7 b (LAMC tvar ckind trm)" - and a31: "\trm ty b. \\c. P7 c trm; \c. P3 c ty\ \ P7 b (APP trm ty)" + and a31: "\trm ty b. \\c. P7 c trm; \c. P3 c ty\ \ P7 b (AppT trm ty)" + and a31a:"\trm co b. \\c. P7 c trm; \c. P5 c co\ \ P7 b (AppC trm co)" and a32: "\var ty trm b. \\c. P3 c ty; \c. P7 c trm; atom var \ b\ \ P7 b (Lam var ty trm)" and a33: "\trm1 trm2 b. \\c. P7 c trm1; \c. P7 c trm2\ \ P7 b (App trm1 trm2)" and a34: "\var ty trm1 trm2 b. \\c. P3 c ty; \c. P7 c trm1; \c. P7 c trm2; atom var \ b\ @@ -369,7 +375,7 @@ and a38: "\pat trm assoc_lst b. \\c. P9 c pat; \c. P7 c trm; \c. P8 c assoc_lst; set (bv (pat)) \* b\ \ P8 b (ACons pat trm assoc_lst)" and a39: "\string tvars cvars vars b. \\c. P11 c tvars; \c. P12 c cvars; \c. P10 c vars\ - \ P9 b (K string tvars cvars vars)" + \ P9 b (Kpat string tvars cvars vars)" and a40: "\b. P10 b VsNil" and a41: "\var ty vars b. \\c. P3 c ty; \c. P10 c vars\ \ P10 b (VsCons var ty vars)" and a42: "\b. P11 b TvsNil" @@ -437,17 +443,17 @@ apply blast (* GOAL2 *) - apply(subgoal_tac "\pa. ((pa \ (atom (p \ tvar))) \ e \ - supp (Abs (p \ {atom tvar}) (p \ co)) \* pa)") + apply(subgoal_tac "\pa. ((pa \ (atom (p \ cvar))) \ e \ + supp (Abs (p \ {atom cvar}) (p \ co)) \* pa)") apply clarify apply (simp only: perm) - apply(rule_tac t="CAll (p \ tvar) (p \ ckind) (p \ co)" - and s="CAll (pa \ p \ tvar) (p \ ckind) (pa \ p \ co)" in subst) + apply(rule_tac t="CAll (p \ cvar) (p \ ckind) (p \ co)" + and s="CAll (pa \ p \ cvar) (p \ ckind) (pa \ p \ co)" in subst) apply (simp only: eq_iff) apply (rule_tac x="-pa" in exI) apply (simp add: alphas eqvts eqvts_raw supp_abs fv_supp) - apply (rule_tac t="supp (pa \ p \ co) - {atom (pa \ p \ tvar)}" - and s="pa \ (p \ supp co - {p \ atom tvar})" in subst) + apply (rule_tac t="supp (pa \ p \ co) - {atom (pa \ p \ cvar)}" + and s="pa \ (p \ supp co - {p \ atom cvar})" in subst) apply (simp add: eqvts) apply (simp add: eqvts[symmetric]) apply (rule conjI) @@ -518,17 +524,17 @@ apply blast (* GOAL4 a copy-and-paste *) - apply(subgoal_tac "\pa. ((pa \ (atom (p \ tvar))) \ g \ - supp (Abs (p \ {atom tvar}) (p \ trm)) \* pa)") + apply(subgoal_tac "\pa. ((pa \ (atom (p \ cvar))) \ g \ + supp (Abs (p \ {atom cvar}) (p \ trm)) \* pa)") apply clarify apply (simp only: perm) - apply(rule_tac t="LAMC (p \ tvar) (p \ ckind) (p \ trm)" - and s="LAMC (pa \ p \ tvar) (p \ ckind) (pa \ p \ trm)" in subst) + apply(rule_tac t="LAMC (p \ cvar) (p \ ckind) (p \ trm)" + and s="LAMC (pa \ p \ cvar) (p \ ckind) (pa \ p \ trm)" in subst) apply (simp only: eq_iff) apply (rule_tac x="-pa" in exI) apply (simp add: alphas eqvts eqvts_raw supp_abs fv_supp) - apply (rule_tac t="supp (pa \ p \ trm) - {atom (pa \ p \ tvar)}" - and s="pa \ (p \ supp trm - {p \ atom tvar})" in subst) + apply (rule_tac t="supp (pa \ p \ trm) - {atom (pa \ p \ cvar)}" + and s="pa \ (p \ supp trm - {p \ atom cvar})" in subst) apply (simp add: eqvts) apply (simp add: eqvts[symmetric]) apply (rule conjI)