added an FCB for res (will not define evry function, but is a good datapoint)
authorChristian Urban <urbanc@in.tum.de>
Wed, 04 Jan 2012 17:42:16 +0000
changeset 3105 1b0d230445ce
parent 3104 f7c4b8e6918b
child 3107 e26e933efba0
added an FCB for res (will not define evry function, but is a good datapoint)
Nominal/Nominal2_FCB.thy
--- a/Nominal/Nominal2_FCB.thy	Tue Jan 03 11:43:27 2012 +0000
+++ b/Nominal/Nominal2_FCB.thy	Wed Jan 04 17:42:16 2012 +0000
@@ -212,67 +212,107 @@
 qed
 
 
-text {* NOT DONE 
 lemma Abs_res_fcb2:
   fixes as bs :: "atom set"
     and x y :: "'b :: fs"
     and c::"'c::fs"
   assumes eq: "[as]res. x = [bs]res. y"
   and fin: "finite as" "finite bs"
-  and fcb1: "as \<sharp>* f as x c"
+  and fcb1: "(as \<inter> supp x) \<sharp>* f (as \<inter> supp x) x c"
   and fresh1: "as \<sharp>* c"
   and fresh2: "bs \<sharp>* c"
-  and perm1: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f as x c) = f (p \<bullet> as) (p \<bullet> x) c"
-  and perm2: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f bs y c) = f (p \<bullet> bs) (p \<bullet> y) c"
-  shows "f as x c = f bs y c"
+  and perm1: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f (as \<inter> supp x) x c) = f (p \<bullet> (as \<inter> supp x)) (p \<bullet> x) c"
+  and perm2: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f (bs \<inter> supp y) y c) = f (p \<bullet> (bs \<inter> supp y)) (p \<bullet> y) c"
+  shows "f (as \<inter> supp x) x c = f (bs \<inter> supp y) y c"
 proof -
-  have "supp (as, x, c) supports (f as x c)"
+  have "supp (as, x, c) supports (f (as \<inter> supp x) x c)"
     unfolding  supports_def fresh_def[symmetric]
-    by (simp add: fresh_Pair perm1 fresh_star_def supp_swap swap_fresh_fresh)
-  then have fin1: "finite (supp (f as x c))"
+    by (simp add: fresh_Pair perm1 fresh_star_def supp_swap swap_fresh_fresh inter_eqvt supp_eqvt)
+  then have fin1: "finite (supp (f (as \<inter> supp x) x c))"
     using fin by (auto intro: supports_finite simp add: finite_supp supp_of_finite_sets supp_Pair)
-  have "supp (bs, y, c) supports (f bs y c)"
+  have "supp (bs, y, c) supports (f (bs \<inter> supp y) y c)"
     unfolding  supports_def fresh_def[symmetric]
-    by (simp add: fresh_Pair perm2 fresh_star_def supp_swap swap_fresh_fresh)
-  then have fin2: "finite (supp (f bs y c))"
+    by (simp add: fresh_Pair perm2 fresh_star_def supp_swap swap_fresh_fresh inter_eqvt supp_eqvt)
+  then have fin2: "finite (supp (f (bs \<inter> supp y) y c))"
     using fin by (auto intro: supports_finite simp add: finite_supp supp_of_finite_sets supp_Pair)
   obtain q::"perm" where 
-    fr1: "(q \<bullet> as) \<sharp>* (x, c, f as x c, f bs y c)" and 
-    fr2: "supp q \<sharp>* ([as]res. x)" and 
-    inc: "supp q \<subseteq> as \<union> (q \<bullet> as)"
-    using at_set_avoiding3[where xs="as" and c="(x, c, f as x c, f bs y c)" and x="[as]res. x"]  
+    fr1: "(q \<bullet> (as \<inter> supp x)) \<sharp>* (x, c, f (as \<inter> supp x) x c, f (bs \<inter> supp y) y c)" and 
+    fr2: "supp q \<sharp>* ([as \<inter> supp x]set. x)" and 
+    inc: "supp q \<subseteq> (as \<inter> supp x) \<union> (q \<bullet> (as \<inter> supp x))"
+    using at_set_avoiding3[where xs="as \<inter> supp x" and c="(x, c, f (as \<inter> supp x) x c, f (bs \<inter> supp y) y c)" 
+      and x="[as \<inter> supp x]set. x"]  
       fin1 fin2 fin
-    by (auto simp add: supp_Pair finite_supp Abs_fresh_star dest: fresh_star_supp_conv)
-  have "[q \<bullet> as]res. (q \<bullet> x) = q \<bullet> ([as]res. x)" by simp
-  also have "\<dots> = [as]res. x"
+    apply (auto simp add: supp_Pair finite_supp Abs_fresh_star dest: fresh_star_supp_conv)
+    done
+  have "[q \<bullet> (as \<inter> supp x)]set. (q \<bullet> x) = q \<bullet> ([as \<inter> supp x]set. x)" by simp
+  also have "\<dots> = [as \<inter> supp x]set. x"
     by (simp only: fr2 perm_supp_eq)
-  finally have "[q \<bullet> as]res. (q \<bullet> x) = [bs]res. y" using eq by simp
+  finally have "[q \<bullet> (as \<inter> supp x)]set. (q \<bullet> x) = [bs \<inter> supp y]set. y" using eq 
+    by(simp add: Abs_eq_res_set)
   then obtain r::perm where 
     qq1: "q \<bullet> x = r \<bullet> y" and 
     qq2: "(q \<bullet> as \<inter> supp (q \<bullet> x)) = r \<bullet> (bs \<inter> supp y)" and 
-    qq3: "supp r \<subseteq> bs \<inter> supp y \<union> q \<bullet> as \<inter> supp (q \<bullet> x)"
+    qq3: "supp r \<subseteq> (bs \<inter> supp y) \<union> q \<bullet> (as \<inter> supp x)"
     apply(drule_tac sym)
-    apply(subst(asm) Abs_eq_res_set)
     apply(simp only: Abs_eq_iff2 alphas)
     apply(erule exE)
     apply(erule conjE)+
     apply(drule_tac x="p" in meta_spec)
-    apply(simp add: set_eqvt)
+    apply(simp add: set_eqvt inter_eqvt supp_eqvt)
     done
-  have "(as \<inter> supp x) \<sharp>* f (as \<inter> supp x) x c" sorry (* FCB? *)
+  have "(as \<inter> supp x) \<sharp>* f (as \<inter> supp x) x c" by (rule fcb1)
   then have "q \<bullet> ((as \<inter> supp x) \<sharp>* f (as \<inter> supp x) x c)"
     by (simp add: permute_bool_def)
   then have "(q \<bullet> (as \<inter> supp x)) \<sharp>* f (q \<bullet> (as \<inter> supp x)) (q \<bullet> x) c"
     apply(simp add: fresh_star_eqvt set_eqvt)
-    sorry (* perm? *)
-  then have "r \<bullet> (bs \<inter> supp y) \<sharp>* f (r \<bullet> (bs \<inter> supp y)) (r \<bullet> y) c" using qq2 
-    apply (simp add: inter_eqvt)
-    sorry
-  (* rest similar reversing it other way around... *)
-  show ?thesis sorry
+    apply(subst (asm) perm1)
+    using inc fresh1 fr1
+    apply(auto simp add: fresh_star_def fresh_Pair)
+    done
+  then have "(r \<bullet> (bs \<inter> supp y)) \<sharp>* f (r \<bullet> (bs \<inter> supp y)) (r \<bullet> y) c" using qq1 qq2 
+    apply(perm_simp)
+    apply simp
+    done
+  then have "r \<bullet> ((bs \<inter> supp y) \<sharp>* f (bs \<inter> supp y) y c)"
+    apply(simp add: fresh_star_eqvt set_eqvt)
+    apply(subst (asm) perm2[symmetric])
+    using qq3 fresh2 fr1
+    apply(auto simp add: set_eqvt fresh_star_def fresh_Pair)
+    done
+  then have fcb2: "(bs \<inter> supp y) \<sharp>* f (bs \<inter> supp y) y c" by (simp add: permute_bool_def)
+  have "f (as \<inter> supp x) x c = q \<bullet> (f (as \<inter> supp x) x c)"
+    apply(rule perm_supp_eq[symmetric])
+    using inc fcb1 fr1 
+    apply (auto simp add: fresh_star_def)
+    done
+  also have "\<dots> = f (q \<bullet> (as \<inter> supp x)) (q \<bullet> x) c" 
+    apply(rule perm1)
+    using inc fresh1 fr1 by (auto simp add: fresh_star_def)
+  also have "\<dots> = f (r \<bullet> (bs \<inter> supp y)) (r \<bullet> y) c" using qq1 qq2 
+    apply(perm_simp)
+    apply simp
+    done
+  also have "\<dots> = r \<bullet> (f (bs \<inter> supp y) y c)"
+    apply(rule perm2[symmetric])
+    using qq3 fresh2 fr1 by (auto simp add: fresh_star_def)
+  also have "... = f (bs \<inter> supp y) y c"
+    apply(rule perm_supp_eq)
+    using qq3 fr1 fcb2 by (auto simp add: fresh_star_def)
+  finally show ?thesis by simp
 qed
-*}
 
+typedef ('a::fs, 'b::fs) ffun = "{f::'a => 'b. finite (supp f)}"
+apply(subgoal_tac "\<exists>x::'b::fs. x \<in> (UNIV::('b::fs) set)")
+apply(erule exE)
+apply(rule_tac x="\<lambda>_. x" in exI)
+apply(auto)
+apply(rule_tac S="supp x" in supports_finite)
+apply(simp add: supports_def)
+apply(perm_simp)
+apply(simp add: fresh_def[symmetric])
+apply(simp add: swap_fresh_fresh)
+apply(simp add: finite_supp)
+done
 
 lemma Abs_lst_fcb2:
   fixes as bs :: "atom list"