diff -r f8029d8c88a9 -r 30550327651a Quot/Nominal/LFex.thy --- a/Quot/Nominal/LFex.thy Mon Feb 08 10:47:19 2010 +0100 +++ b/Quot/Nominal/LFex.thy Mon Feb 08 11:41:25 2010 +0100 @@ -403,9 +403,11 @@ as "permute :: perm \ trm \ trm" +lemmas permute_ktt[simp] = permute_kind_permute_ty_permute_trm.simps[quot_lifted] + lemma perm_zero_ok: "0 \ (x :: KIND) = x \ 0 \ (y :: TY) = y \ 0 \ (z :: TRM) = z" apply (induct rule: KIND_TY_TRM_induct) -apply (simp_all add: permute_kind_permute_ty_permute_trm.simps[quot_lifted]) +apply (simp_all) done lemma perm_add_ok: @@ -413,7 +415,7 @@ "((p + q) \ (x2 :: TY) = p \ q \ x2)" "((p + q) \ (x3 :: TRM) = p \ q \ x3)" apply (induct x1 and x2 and x3 rule: KIND_TY_TRM_inducts) -apply (simp_all add: permute_kind_permute_ty_permute_trm.simps[quot_lifted]) +apply (simp_all) done instance @@ -423,8 +425,6 @@ end -lemmas permute_ktt[simp] = permute_kind_permute_ty_permute_trm.simps[quot_lifted] - lemmas AKIND_ATY_ATRM_inducts = akind_aty_atrm.inducts[unfolded alpha_gen, quot_lifted, folded alpha_gen] lemmas KIND_TY_TRM_INJECT = akind_aty_atrm_inj[unfolded alpha_gen, quot_lifted, folded alpha_gen]