# HG changeset patch # User Christian Urban # Date 1265204208 -3600 # Node ID 68bdd55236083330c485e211ac71892c2952e1af # Parent b1ca92ea3a86e5cdf55339841b3ccbae3ecd4776# Parent c1b469325033f1042f6508148e36414b83d7a4a3 merged diff -r c1b469325033 -r 68bdd5523608 Quot/Nominal/Terms.thy --- a/Quot/Nominal/Terms.thy Wed Feb 03 14:28:00 2010 +0100 +++ b/Quot/Nominal/Terms.thy Wed Feb 03 14:36:48 2010 +0100 @@ -822,6 +822,18 @@ "(alphalts ===> op =) rbv5 rbv5" sorry +lemma bv_list_rsp: + shows "t \5 s \ True" + and "x \l y \ rbv5 x = rbv5 y" +apply(induct rule: alpha5_alphalts.inducts) +apply(simp_all) +done + +lemma + shows "(alphalts ===> op =) rbv5 rbv5" + by (simp add: bv_list_rsp) + + instantiation trm5 and lts :: pt begin