Nominal/Nominal2_Base_Exec.thy
changeset 3179 9eeea01bdbc0
parent 3176 31372760c2fb
child 3182 5335c0ea743a
equal deleted inserted replaced
3178:a331468b2f5a 3179:9eeea01bdbc0
  1009 lemma finite_eqvt [eqvt]:
  1009 lemma finite_eqvt [eqvt]:
  1010   shows "p \<bullet> finite A = finite (p \<bullet> A)"
  1010   shows "p \<bullet> finite A = finite (p \<bullet> A)"
  1011 unfolding finite_def
  1011 unfolding finite_def
  1012 by (perm_simp) (rule refl)
  1012 by (perm_simp) (rule refl)
  1013 
  1013 
  1014 lemma Let_eqvt [eqvt]:
       
  1015   "p \<bullet> Let x y = Let (p \<bullet> x) (p \<bullet> y)"
       
  1016   unfolding Let_def permute_fun_app_eq ..
       
  1017 
  1014 
  1018 subsubsection {* Equivariance for product operations *}
  1015 subsubsection {* Equivariance for product operations *}
  1019 
  1016 
  1020 lemma fst_eqvt [eqvt]:
  1017 lemma fst_eqvt [eqvt]:
  1021   shows "p \<bullet> (fst x) = fst (p \<bullet> x)"
  1018   shows "p \<bullet> (fst x) = fst (p \<bullet> x)"
  1672   using a
  1669   using a
  1673   unfolding eqvt_def
  1670   unfolding eqvt_def
  1674   unfolding supp_def 
  1671   unfolding supp_def 
  1675   by simp
  1672   by simp
  1676 
  1673 
       
  1674 lemma fresh_fun_eqvt:
       
  1675   assumes a: "eqvt f"
       
  1676   shows "a \<sharp> f"
       
  1677   using a
       
  1678   unfolding fresh_def
       
  1679   by (simp add: supp_fun_eqvt)
       
  1680 
  1677 lemma fresh_fun_eqvt_app:
  1681 lemma fresh_fun_eqvt_app:
  1678   assumes a: "eqvt f"
  1682   assumes a: "eqvt f"
  1679   shows "a \<sharp> x \<Longrightarrow> a \<sharp> f x"
  1683   shows "a \<sharp> x \<Longrightarrow> a \<sharp> f x"
  1680 proof -
  1684 proof -
  1681   from a have "supp f = {}" by (simp add: supp_fun_eqvt)
  1685   from a have "supp f = {}" by (simp add: supp_fun_eqvt)
  1720 using fresh
  1724 using fresh
  1721 unfolding fresh_def
  1725 unfolding fresh_def
  1722 using supp_eqvt_at[OF asm fin]
  1726 using supp_eqvt_at[OF asm fin]
  1723 by auto
  1727 by auto
  1724 
  1728 
       
  1729 text {* for handling of freshness of functions *}
       
  1730 
       
  1731 simproc_setup fresh_fun_simproc ("a \<sharp> (f::'a::pt \<Rightarrow>'b::pt)") = {* fn _ => fn ss => fn ctrm =>
       
  1732   let 
       
  1733     val Const(@{const_name fresh}, _) $ _ $ f = term_of ctrm
       
  1734   in
       
  1735     case (Term.add_frees f [], Term.add_vars f []) of
       
  1736       ([], []) => SOME(@{thm fresh_fun_eqvt[simplified eqvt_def, THEN Eq_TrueI]})
       
  1737     | (x::_, []) => let
       
  1738          val thy = Proof_Context.theory_of (Simplifier.the_context ss)        
       
  1739          val argx = Free x
       
  1740          val absf = absfree x f
       
  1741          val cty_inst = [SOME (ctyp_of thy (fastype_of argx)), SOME (ctyp_of thy (fastype_of f))]
       
  1742          val ctrm_inst = [NONE, SOME (cterm_of thy absf), SOME (cterm_of thy argx)] 
       
  1743          val thm = Drule.instantiate' cty_inst ctrm_inst @{thm fresh_fun_app}
       
  1744       in
       
  1745         SOME(thm RS @{thm Eq_TrueI}) 
       
  1746       end  
       
  1747     | (_, _) => NONE
       
  1748   end
       
  1749 *}
  1725 
  1750 
  1726 subsection {* helper functions for nominal_functions *}
  1751 subsection {* helper functions for nominal_functions *}
  1727 
  1752 
  1728 lemma THE_defaultI2:
  1753 lemma THE_defaultI2:
  1729   assumes "\<exists>!x. P x" "\<And>x. P x \<Longrightarrow> Q x"
  1754   assumes "\<exists>!x. P x" "\<And>x. P x \<Longrightarrow> Q x"
  3074 qed
  3099 qed
  3075 
  3100 
  3076 text {* packaging the freshness lemma into a function *}
  3101 text {* packaging the freshness lemma into a function *}
  3077 
  3102 
  3078 definition
  3103 definition
  3079   fresh_fun :: "('a::at \<Rightarrow> 'b::pt) \<Rightarrow> 'b"
  3104   Fresh :: "('a::at \<Rightarrow> 'b::pt) \<Rightarrow> 'b"
  3080 where
  3105 where
  3081   "fresh_fun h = (THE x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x)"
  3106   "Fresh h = (THE x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x)"
  3082 
  3107 
  3083 lemma fresh_fun_apply:
  3108 lemma Fresh_apply:
  3084   fixes h :: "'a::at \<Rightarrow> 'b::pt"
  3109   fixes h :: "'a::at \<Rightarrow> 'b::pt"
  3085   assumes a: "\<exists>a. atom a \<sharp> (h, h a)"
  3110   assumes a: "\<exists>a. atom a \<sharp> (h, h a)"
  3086   assumes b: "atom a \<sharp> h"
  3111   assumes b: "atom a \<sharp> h"
  3087   shows "fresh_fun h = h a"
  3112   shows "Fresh h = h a"
  3088 unfolding fresh_fun_def
  3113 unfolding Fresh_def
  3089 proof (rule the_equality)
  3114 proof (rule the_equality)
  3090   show "\<forall>a'. atom a' \<sharp> h \<longrightarrow> h a' = h a"
  3115   show "\<forall>a'. atom a' \<sharp> h \<longrightarrow> h a' = h a"
  3091   proof (intro strip)
  3116   proof (intro strip)
  3092     fix a':: 'a
  3117     fix a':: 'a
  3093     assume c: "atom a' \<sharp> h"
  3118     assume c: "atom a' \<sharp> h"
  3098   fix fr :: 'b
  3123   fix fr :: 'b
  3099   assume "\<forall>a. atom a \<sharp> h \<longrightarrow> h a = fr"
  3124   assume "\<forall>a. atom a \<sharp> h \<longrightarrow> h a = fr"
  3100   with b show "fr = h a" by auto
  3125   with b show "fr = h a" by auto
  3101 qed
  3126 qed
  3102 
  3127 
  3103 lemma fresh_fun_apply':
  3128 lemma Fresh_apply':
  3104   fixes h :: "'a::at \<Rightarrow> 'b::pt"
  3129   fixes h :: "'a::at \<Rightarrow> 'b::pt"
  3105   assumes a: "atom a \<sharp> h" "atom a \<sharp> h a"
  3130   assumes a: "atom a \<sharp> h" "atom a \<sharp> h a"
  3106   shows "fresh_fun h = h a"
  3131   shows "Fresh h = h a"
  3107   apply (rule fresh_fun_apply)
  3132   apply (rule Fresh_apply)
  3108   apply (auto simp add: fresh_Pair intro: a)
  3133   apply (auto simp add: fresh_Pair intro: a)
  3109   done
  3134   done
  3110 
  3135 
  3111 lemma fresh_fun_eqvt:
  3136 lemma Fresh_eqvt:
  3112   fixes h :: "'a::at \<Rightarrow> 'b::pt"
  3137   fixes h :: "'a::at \<Rightarrow> 'b::pt"
  3113   assumes a: "\<exists>a. atom a \<sharp> (h, h a)"
  3138   assumes a: "\<exists>a. atom a \<sharp> (h, h a)"
  3114   shows "p \<bullet> (fresh_fun h) = fresh_fun (p \<bullet> h)"
  3139   shows "p \<bullet> (Fresh h) = Fresh (p \<bullet> h)"
  3115   using a
  3140   using a
  3116   apply (clarsimp simp add: fresh_Pair)
  3141   apply (clarsimp simp add: fresh_Pair)
  3117   apply (subst fresh_fun_apply', assumption+)
  3142   apply (subst Fresh_apply', assumption+)
  3118   apply (drule fresh_permute_iff [where p=p, THEN iffD2])
  3143   apply (drule fresh_permute_iff [where p=p, THEN iffD2])
  3119   apply (drule fresh_permute_iff [where p=p, THEN iffD2])
  3144   apply (drule fresh_permute_iff [where p=p, THEN iffD2])
  3120   apply (simp only: atom_eqvt permute_fun_app_eq [where f=h])
  3145   apply (simp only: atom_eqvt permute_fun_app_eq [where f=h])
  3121   apply (erule (1) fresh_fun_apply' [symmetric])
  3146   apply (erule (1) Fresh_apply' [symmetric])
  3122   done
  3147   done
  3123 
  3148 
  3124 lemma fresh_fun_supports:
  3149 lemma Fresh_supports:
  3125   fixes h :: "'a::at \<Rightarrow> 'b::pt"
  3150   fixes h :: "'a::at \<Rightarrow> 'b::pt"
  3126   assumes a: "\<exists>a. atom a \<sharp> (h, h a)"
  3151   assumes a: "\<exists>a. atom a \<sharp> (h, h a)"
  3127   shows "(supp h) supports (fresh_fun h)"
  3152   shows "(supp h) supports (Fresh h)"
  3128   apply (simp add: supports_def fresh_def [symmetric])
  3153   apply (simp add: supports_def fresh_def [symmetric])
  3129   apply (simp add: fresh_fun_eqvt [OF a] swap_fresh_fresh)
  3154   apply (simp add: Fresh_eqvt [OF a] swap_fresh_fresh)
  3130   done
  3155   done
  3131 
  3156 
  3132 notation fresh_fun (binder "FRESH " 10)
  3157 notation Fresh (binder "FRESH " 10)
  3133 
  3158 
  3134 lemma FRESH_f_iff:
  3159 lemma FRESH_f_iff:
  3135   fixes P :: "'a::at \<Rightarrow> 'b::pure"
  3160   fixes P :: "'a::at \<Rightarrow> 'b::pure"
  3136   fixes f :: "'b \<Rightarrow> 'c::pure"
  3161   fixes f :: "'b \<Rightarrow> 'c::pure"
  3137   assumes P: "finite (supp P)"
  3162   assumes P: "finite (supp P)"
  3138   shows "(FRESH x. f (P x)) = f (FRESH x. P x)"
  3163   shows "(FRESH x. f (P x)) = f (FRESH x. P x)"
  3139 proof -
  3164 proof -
  3140   obtain a::'a where "atom a \<sharp> P" using P by (rule obtain_fresh')
  3165   obtain a::'a where "atom a \<sharp> P" using P by (rule obtain_fresh')
  3141   show "(FRESH x. f (P x)) = f (FRESH x. P x)"
  3166   show "(FRESH x. f (P x)) = f (FRESH x. P x)"
  3142     apply (subst fresh_fun_apply' [where a=a, OF _ pure_fresh])
  3167     apply (subst Fresh_apply' [where a=a, OF _ pure_fresh])
  3143     apply (cut_tac `atom a \<sharp> P`)
  3168     apply (cut_tac `atom a \<sharp> P`)
  3144     apply (simp add: fresh_conv_MOST)
  3169     apply (simp add: fresh_conv_MOST)
  3145     apply (elim MOST_rev_mp, rule MOST_I, clarify)
  3170     apply (elim MOST_rev_mp, rule MOST_I, clarify)
  3146     apply (simp add: permute_fun_def permute_pure fun_eq_iff)
  3171     apply (simp add: permute_fun_def permute_pure fun_eq_iff)
  3147     apply (subst fresh_fun_apply' [where a=a, OF `atom a \<sharp> P` pure_fresh])
  3172     apply (subst Fresh_apply' [where a=a, OF `atom a \<sharp> P` pure_fresh])
  3148     apply (rule refl)
  3173     apply (rule refl)
  3149     done
  3174     done
  3150 qed
  3175 qed
  3151 
  3176 
  3152 lemma FRESH_binop_iff:
  3177 lemma FRESH_binop_iff:
  3159 proof -
  3184 proof -
  3160   from assms have "finite (supp (P, Q))" by (simp add: supp_Pair)
  3185   from assms have "finite (supp (P, Q))" by (simp add: supp_Pair)
  3161   then obtain a::'a where "atom a \<sharp> (P, Q)" by (rule obtain_fresh') 
  3186   then obtain a::'a where "atom a \<sharp> (P, Q)" by (rule obtain_fresh') 
  3162   then have "atom a \<sharp> P" and "atom a \<sharp> Q" by (simp_all add: fresh_Pair)
  3187   then have "atom a \<sharp> P" and "atom a \<sharp> Q" by (simp_all add: fresh_Pair)
  3163   show ?thesis
  3188   show ?thesis
  3164     apply (subst fresh_fun_apply' [where a=a, OF _ pure_fresh])
  3189     apply (subst Fresh_apply' [where a=a, OF _ pure_fresh])
  3165     apply (cut_tac `atom a \<sharp> P` `atom a \<sharp> Q`)
  3190     apply (cut_tac `atom a \<sharp> P` `atom a \<sharp> Q`)
  3166     apply (simp add: fresh_conv_MOST)
  3191     apply (simp add: fresh_conv_MOST)
  3167     apply (elim MOST_rev_mp, rule MOST_I, clarify)
  3192     apply (elim MOST_rev_mp, rule MOST_I, clarify)
  3168     apply (simp add: permute_fun_def permute_pure fun_eq_iff)
  3193     apply (simp add: permute_fun_def permute_pure fun_eq_iff)
  3169     apply (subst fresh_fun_apply' [where a=a, OF `atom a \<sharp> P` pure_fresh])
  3194     apply (subst Fresh_apply' [where a=a, OF `atom a \<sharp> P` pure_fresh])
  3170     apply (subst fresh_fun_apply' [where a=a, OF `atom a \<sharp> Q` pure_fresh])
  3195     apply (subst Fresh_apply' [where a=a, OF `atom a \<sharp> Q` pure_fresh])
  3171     apply (rule refl)
  3196     apply (rule refl)
  3172     done
  3197     done
  3173 qed
  3198 qed
  3174 
  3199 
  3175 lemma FRESH_conj_iff:
  3200 lemma FRESH_conj_iff: