--- 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 \<Rightarrow> trm \<Rightarrow> trm"
+lemmas permute_ktt[simp] = permute_kind_permute_ty_permute_trm.simps[quot_lifted]
+
lemma perm_zero_ok: "0 \<bullet> (x :: KIND) = x \<and> 0 \<bullet> (y :: TY) = y \<and> 0 \<bullet> (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) \<bullet> (x2 :: TY) = p \<bullet> q \<bullet> x2)"
"((p + q) \<bullet> (x3 :: TRM) = p \<bullet> q \<bullet> 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]