1724 using fresh |
1728 using fresh |
1725 unfolding fresh_def |
1729 unfolding fresh_def |
1726 using supp_eqvt_at[OF asm fin] |
1730 using supp_eqvt_at[OF asm fin] |
1727 by auto |
1731 by auto |
1728 |
1732 |
|
1733 text {* for handling of freshness of functions *} |
|
1734 |
|
1735 simproc_setup fresh_fun_simproc ("a \<sharp> (f::'a::pt \<Rightarrow>'b::pt)") = {* fn _ => fn ss => fn ctrm => |
|
1736 let |
|
1737 val Const(@{const_name fresh}, _) $ _ $ f = term_of ctrm |
|
1738 in |
|
1739 case (Term.add_frees f [], Term.add_vars f []) of |
|
1740 ([], []) => SOME(@{thm fresh_fun_eqvt[simplified eqvt_def, THEN Eq_TrueI]}) |
|
1741 | (x::_, []) => let |
|
1742 val thy = Proof_Context.theory_of (Simplifier.the_context ss) |
|
1743 val argx = Free x |
|
1744 val absf = absfree x f |
|
1745 val cty_inst = [SOME (ctyp_of thy (fastype_of argx)), SOME (ctyp_of thy (fastype_of f))] |
|
1746 val ctrm_inst = [NONE, SOME (cterm_of thy absf), SOME (cterm_of thy argx)] |
|
1747 val thm = Drule.instantiate' cty_inst ctrm_inst @{thm fresh_fun_app} |
|
1748 in |
|
1749 SOME(thm RS @{thm Eq_TrueI}) |
|
1750 end |
|
1751 | (_, _) => NONE |
|
1752 end |
|
1753 *} |
1729 |
1754 |
1730 subsection {* helper functions for nominal_functions *} |
1755 subsection {* helper functions for nominal_functions *} |
1731 |
1756 |
1732 lemma THE_defaultI2: |
1757 lemma THE_defaultI2: |
1733 assumes "\<exists>!x. P x" "\<And>x. P x \<Longrightarrow> Q x" |
1758 assumes "\<exists>!x. P x" "\<And>x. P x \<Longrightarrow> Q x" |
3078 qed |
3103 qed |
3079 |
3104 |
3080 text {* packaging the freshness lemma into a function *} |
3105 text {* packaging the freshness lemma into a function *} |
3081 |
3106 |
3082 definition |
3107 definition |
3083 fresh_fun :: "('a::at \<Rightarrow> 'b::pt) \<Rightarrow> 'b" |
3108 Fresh :: "('a::at \<Rightarrow> 'b::pt) \<Rightarrow> 'b" |
3084 where |
3109 where |
3085 "fresh_fun h = (THE x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x)" |
3110 "Fresh h = (THE x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x)" |
3086 |
3111 |
3087 lemma fresh_fun_apply: |
3112 lemma Fresh_apply: |
3088 fixes h :: "'a::at \<Rightarrow> 'b::pt" |
3113 fixes h :: "'a::at \<Rightarrow> 'b::pt" |
3089 assumes a: "\<exists>a. atom a \<sharp> (h, h a)" |
3114 assumes a: "\<exists>a. atom a \<sharp> (h, h a)" |
3090 assumes b: "atom a \<sharp> h" |
3115 assumes b: "atom a \<sharp> h" |
3091 shows "fresh_fun h = h a" |
3116 shows "Fresh h = h a" |
3092 unfolding fresh_fun_def |
3117 unfolding Fresh_def |
3093 proof (rule the_equality) |
3118 proof (rule the_equality) |
3094 show "\<forall>a'. atom a' \<sharp> h \<longrightarrow> h a' = h a" |
3119 show "\<forall>a'. atom a' \<sharp> h \<longrightarrow> h a' = h a" |
3095 proof (intro strip) |
3120 proof (intro strip) |
3096 fix a':: 'a |
3121 fix a':: 'a |
3097 assume c: "atom a' \<sharp> h" |
3122 assume c: "atom a' \<sharp> h" |
3102 fix fr :: 'b |
3127 fix fr :: 'b |
3103 assume "\<forall>a. atom a \<sharp> h \<longrightarrow> h a = fr" |
3128 assume "\<forall>a. atom a \<sharp> h \<longrightarrow> h a = fr" |
3104 with b show "fr = h a" by auto |
3129 with b show "fr = h a" by auto |
3105 qed |
3130 qed |
3106 |
3131 |
3107 lemma fresh_fun_apply': |
3132 lemma Fresh_apply': |
3108 fixes h :: "'a::at \<Rightarrow> 'b::pt" |
3133 fixes h :: "'a::at \<Rightarrow> 'b::pt" |
3109 assumes a: "atom a \<sharp> h" "atom a \<sharp> h a" |
3134 assumes a: "atom a \<sharp> h" "atom a \<sharp> h a" |
3110 shows "fresh_fun h = h a" |
3135 shows "Fresh h = h a" |
3111 apply (rule fresh_fun_apply) |
3136 apply (rule Fresh_apply) |
3112 apply (auto simp add: fresh_Pair intro: a) |
3137 apply (auto simp add: fresh_Pair intro: a) |
3113 done |
3138 done |
3114 |
3139 |
3115 lemma fresh_fun_eqvt: |
3140 lemma Fresh_eqvt: |
3116 fixes h :: "'a::at \<Rightarrow> 'b::pt" |
3141 fixes h :: "'a::at \<Rightarrow> 'b::pt" |
3117 assumes a: "\<exists>a. atom a \<sharp> (h, h a)" |
3142 assumes a: "\<exists>a. atom a \<sharp> (h, h a)" |
3118 shows "p \<bullet> (fresh_fun h) = fresh_fun (p \<bullet> h)" |
3143 shows "p \<bullet> (Fresh h) = Fresh (p \<bullet> h)" |
3119 using a |
3144 using a |
3120 apply (clarsimp simp add: fresh_Pair) |
3145 apply (clarsimp simp add: fresh_Pair) |
3121 apply (subst fresh_fun_apply', assumption+) |
3146 apply (subst Fresh_apply', assumption+) |
3122 apply (drule fresh_permute_iff [where p=p, THEN iffD2]) |
3147 apply (drule fresh_permute_iff [where p=p, THEN iffD2]) |
3123 apply (drule fresh_permute_iff [where p=p, THEN iffD2]) |
3148 apply (drule fresh_permute_iff [where p=p, THEN iffD2]) |
3124 apply (simp only: atom_eqvt permute_fun_app_eq [where f=h]) |
3149 apply (simp only: atom_eqvt permute_fun_app_eq [where f=h]) |
3125 apply (erule (1) fresh_fun_apply' [symmetric]) |
3150 apply (erule (1) Fresh_apply' [symmetric]) |
3126 done |
3151 done |
3127 |
3152 |
3128 lemma fresh_fun_supports: |
3153 lemma Fresh_supports: |
3129 fixes h :: "'a::at \<Rightarrow> 'b::pt" |
3154 fixes h :: "'a::at \<Rightarrow> 'b::pt" |
3130 assumes a: "\<exists>a. atom a \<sharp> (h, h a)" |
3155 assumes a: "\<exists>a. atom a \<sharp> (h, h a)" |
3131 shows "(supp h) supports (fresh_fun h)" |
3156 shows "(supp h) supports (Fresh h)" |
3132 apply (simp add: supports_def fresh_def [symmetric]) |
3157 apply (simp add: supports_def fresh_def [symmetric]) |
3133 apply (simp add: fresh_fun_eqvt [OF a] swap_fresh_fresh) |
3158 apply (simp add: Fresh_eqvt [OF a] swap_fresh_fresh) |
3134 done |
3159 done |
3135 |
3160 |
3136 notation fresh_fun (binder "FRESH " 10) |
3161 notation Fresh (binder "FRESH " 10) |
3137 |
3162 |
3138 lemma FRESH_f_iff: |
3163 lemma FRESH_f_iff: |
3139 fixes P :: "'a::at \<Rightarrow> 'b::pure" |
3164 fixes P :: "'a::at \<Rightarrow> 'b::pure" |
3140 fixes f :: "'b \<Rightarrow> 'c::pure" |
3165 fixes f :: "'b \<Rightarrow> 'c::pure" |
3141 assumes P: "finite (supp P)" |
3166 assumes P: "finite (supp P)" |
3142 shows "(FRESH x. f (P x)) = f (FRESH x. P x)" |
3167 shows "(FRESH x. f (P x)) = f (FRESH x. P x)" |
3143 proof - |
3168 proof - |
3144 obtain a::'a where "atom a \<sharp> P" using P by (rule obtain_fresh') |
3169 obtain a::'a where "atom a \<sharp> P" using P by (rule obtain_fresh') |
3145 show "(FRESH x. f (P x)) = f (FRESH x. P x)" |
3170 show "(FRESH x. f (P x)) = f (FRESH x. P x)" |
3146 apply (subst fresh_fun_apply' [where a=a, OF _ pure_fresh]) |
3171 apply (subst Fresh_apply' [where a=a, OF _ pure_fresh]) |
3147 apply (cut_tac `atom a \<sharp> P`) |
3172 apply (cut_tac `atom a \<sharp> P`) |
3148 apply (simp add: fresh_conv_MOST) |
3173 apply (simp add: fresh_conv_MOST) |
3149 apply (elim MOST_rev_mp, rule MOST_I, clarify) |
3174 apply (elim MOST_rev_mp, rule MOST_I, clarify) |
3150 apply (simp add: permute_fun_def permute_pure fun_eq_iff) |
3175 apply (simp add: permute_fun_def permute_pure fun_eq_iff) |
3151 apply (subst fresh_fun_apply' [where a=a, OF `atom a \<sharp> P` pure_fresh]) |
3176 apply (subst Fresh_apply' [where a=a, OF `atom a \<sharp> P` pure_fresh]) |
3152 apply (rule refl) |
3177 apply (rule refl) |
3153 done |
3178 done |
3154 qed |
3179 qed |
3155 |
3180 |
3156 lemma FRESH_binop_iff: |
3181 lemma FRESH_binop_iff: |
3163 proof - |
3188 proof - |
3164 from assms have "finite (supp (P, Q))" by (simp add: supp_Pair) |
3189 from assms have "finite (supp (P, Q))" by (simp add: supp_Pair) |
3165 then obtain a::'a where "atom a \<sharp> (P, Q)" by (rule obtain_fresh') |
3190 then obtain a::'a where "atom a \<sharp> (P, Q)" by (rule obtain_fresh') |
3166 then have "atom a \<sharp> P" and "atom a \<sharp> Q" by (simp_all add: fresh_Pair) |
3191 then have "atom a \<sharp> P" and "atom a \<sharp> Q" by (simp_all add: fresh_Pair) |
3167 show ?thesis |
3192 show ?thesis |
3168 apply (subst fresh_fun_apply' [where a=a, OF _ pure_fresh]) |
3193 apply (subst Fresh_apply' [where a=a, OF _ pure_fresh]) |
3169 apply (cut_tac `atom a \<sharp> P` `atom a \<sharp> Q`) |
3194 apply (cut_tac `atom a \<sharp> P` `atom a \<sharp> Q`) |
3170 apply (simp add: fresh_conv_MOST) |
3195 apply (simp add: fresh_conv_MOST) |
3171 apply (elim MOST_rev_mp, rule MOST_I, clarify) |
3196 apply (elim MOST_rev_mp, rule MOST_I, clarify) |
3172 apply (simp add: permute_fun_def permute_pure fun_eq_iff) |
3197 apply (simp add: permute_fun_def permute_pure fun_eq_iff) |
3173 apply (subst fresh_fun_apply' [where a=a, OF `atom a \<sharp> P` pure_fresh]) |
3198 apply (subst Fresh_apply' [where a=a, OF `atom a \<sharp> P` pure_fresh]) |
3174 apply (subst fresh_fun_apply' [where a=a, OF `atom a \<sharp> Q` pure_fresh]) |
3199 apply (subst Fresh_apply' [where a=a, OF `atom a \<sharp> Q` pure_fresh]) |
3175 apply (rule refl) |
3200 apply (rule refl) |
3176 done |
3201 done |
3177 qed |
3202 qed |
3178 |
3203 |
3179 lemma FRESH_conj_iff: |
3204 lemma FRESH_conj_iff: |