proved that bv for lists respects alpha for terms
authorChristian Urban <urbanc@in.tum.de>
Wed, 03 Feb 2010 14:36:22 +0100
changeset 1053 b1ca92ea3a86
parent 1051 277301dc5c4c
child 1054 68bdd5523608
proved that bv for lists respects alpha for terms
Quot/Nominal/Terms.thy
--- a/Quot/Nominal/Terms.thy	Wed Feb 03 14:22:25 2010 +0100
+++ b/Quot/Nominal/Terms.thy	Wed Feb 03 14:36:22 2010 +0100
@@ -822,6 +822,18 @@
 "(alphalts ===> op =) rbv5 rbv5"
 sorry
 
+lemma bv_list_rsp:
+  shows "t \<approx>5 s \<Longrightarrow> True"
+  and "x \<approx>l y \<Longrightarrow> 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