Quot/Nominal/Terms.thy
changeset 1179 789fbba5c23f
parent 1171 62632eec979c
child 1180 3f36936f1280
equal deleted inserted replaced
1178:275a1cb3f2ba 1179:789fbba5c23f
   269 
   269 
   270 lemma trm1_induct_strong:
   270 lemma trm1_induct_strong:
   271   assumes "\<And>name b. P b (Vr1 name)"
   271   assumes "\<And>name b. P b (Vr1 name)"
   272   and     "\<And>rtrm11 rtrm12 b. \<lbrakk>\<And>c. P c rtrm11; \<And>c. P c rtrm12\<rbrakk> \<Longrightarrow> P b (Ap1 rtrm11 rtrm12)"
   272   and     "\<And>rtrm11 rtrm12 b. \<lbrakk>\<And>c. P c rtrm11; \<And>c. P c rtrm12\<rbrakk> \<Longrightarrow> P b (Ap1 rtrm11 rtrm12)"
   273   and     "\<And>name rtrm1 b. \<lbrakk>\<And>c. P c rtrm1; (atom name) \<sharp> b\<rbrakk> \<Longrightarrow> P b (Lm1 name rtrm1)"
   273   and     "\<And>name rtrm1 b. \<lbrakk>\<And>c. P c rtrm1; (atom name) \<sharp> b\<rbrakk> \<Longrightarrow> P b (Lm1 name rtrm1)"
   274   and     "\<And>bp rtrm11 rtrm12 b. \<lbrakk>\<And>c. P c rtrm11; \<And>c. P c rtrm12; bp1 bp \<sharp>* b\<rbrakk> \<Longrightarrow> P b (Lt1 bp rtrm11 rtrm12)"
   274   and     "\<And>bp rtrm11 rtrm12 b. \<lbrakk>\<And>c. P c rtrm11; \<And>c. P c rtrm12; bv1 bp \<sharp>* b\<rbrakk> \<Longrightarrow> P b (Lt1 bp rtrm11 rtrm12)"
   275   shows   "P a rtrma"
   275   shows   "P a rtrma"
   276 sorry
   276 sorry
   277 
   277 
   278 section {*** lets with single assignments ***}
   278 section {*** lets with single assignments ***}
   279 
   279