Nominal-General/Nominal2_Base.thy
changeset 2568 8193bbaa07fe
parent 2567 41137dc935ff
child 2569 94750b31a97d
equal deleted inserted replaced
2567:41137dc935ff 2568:8193bbaa07fe
     1 (*  Title:      Nominal2_Base
       
     2     Authors:    Brian Huffman, Christian Urban
       
     3 
       
     4     Basic definitions and lemma infrastructure for 
       
     5     Nominal Isabelle. 
       
     6 *)
       
     7 theory Nominal2_Base
       
     8 imports Main Infinite_Set
       
     9         "~~/src/HOL/Quotient_Examples/FSet"
       
    10 uses ("nominal_library.ML")
       
    11      ("nominal_atoms.ML")
       
    12 begin
       
    13 
       
    14 section {* Atoms and Sorts *}
       
    15 
       
    16 text {* A simple implementation for atom_sorts is strings. *}
       
    17 (* types atom_sort = string *)
       
    18 
       
    19 text {* To deal with Church-like binding we use trees of  
       
    20   strings as sorts. *}
       
    21 
       
    22 datatype atom_sort = Sort "string" "atom_sort list"
       
    23 
       
    24 datatype atom = Atom atom_sort nat
       
    25 
       
    26 
       
    27 text {* Basic projection function. *}
       
    28 
       
    29 primrec
       
    30   sort_of :: "atom \<Rightarrow> atom_sort"
       
    31 where
       
    32   "sort_of (Atom s i) = s"
       
    33 
       
    34 primrec
       
    35   nat_of :: "atom \<Rightarrow> nat"
       
    36 where
       
    37   "nat_of (Atom s n) = n"
       
    38 
       
    39 
       
    40 text {* There are infinitely many atoms of each sort. *}
       
    41 lemma INFM_sort_of_eq: 
       
    42   shows "INFM a. sort_of a = s"
       
    43 proof -
       
    44   have "INFM i. sort_of (Atom s i) = s" by simp
       
    45   moreover have "inj (Atom s)" by (simp add: inj_on_def)
       
    46   ultimately show "INFM a. sort_of a = s" by (rule INFM_inj)
       
    47 qed
       
    48 
       
    49 lemma infinite_sort_of_eq:
       
    50   shows "infinite {a. sort_of a = s}"
       
    51   using INFM_sort_of_eq unfolding INFM_iff_infinite .
       
    52 
       
    53 lemma atom_infinite [simp]: 
       
    54   shows "infinite (UNIV :: atom set)"
       
    55   using subset_UNIV infinite_sort_of_eq
       
    56   by (rule infinite_super)
       
    57 
       
    58 lemma obtain_atom:
       
    59   fixes X :: "atom set"
       
    60   assumes X: "finite X"
       
    61   obtains a where "a \<notin> X" "sort_of a = s"
       
    62 proof -
       
    63   from X have "MOST a. a \<notin> X"
       
    64     unfolding MOST_iff_cofinite by simp
       
    65   with INFM_sort_of_eq
       
    66   have "INFM a. sort_of a = s \<and> a \<notin> X"
       
    67     by (rule INFM_conjI)
       
    68   then obtain a where "a \<notin> X" "sort_of a = s"
       
    69     by (auto elim: INFM_E)
       
    70   then show ?thesis ..
       
    71 qed
       
    72 
       
    73 lemma atom_components_eq_iff:
       
    74   fixes a b :: atom
       
    75   shows "a = b \<longleftrightarrow> sort_of a = sort_of b \<and> nat_of a = nat_of b"
       
    76   by (induct a, induct b, simp)
       
    77 
       
    78 section {* Sort-Respecting Permutations *}
       
    79 
       
    80 typedef perm =
       
    81   "{f. bij f \<and> finite {a. f a \<noteq> a} \<and> (\<forall>a. sort_of (f a) = sort_of a)}"
       
    82 proof
       
    83   show "id \<in> ?perm" by simp
       
    84 qed
       
    85 
       
    86 lemma permI:
       
    87   assumes "bij f" and "MOST x. f x = x" and "\<And>a. sort_of (f a) = sort_of a"
       
    88   shows "f \<in> perm"
       
    89   using assms unfolding perm_def MOST_iff_cofinite by simp
       
    90 
       
    91 lemma perm_is_bij: "f \<in> perm \<Longrightarrow> bij f"
       
    92   unfolding perm_def by simp
       
    93 
       
    94 lemma perm_is_finite: "f \<in> perm \<Longrightarrow> finite {a. f a \<noteq> a}"
       
    95   unfolding perm_def by simp
       
    96 
       
    97 lemma perm_is_sort_respecting: "f \<in> perm \<Longrightarrow> sort_of (f a) = sort_of a"
       
    98   unfolding perm_def by simp
       
    99 
       
   100 lemma perm_MOST: "f \<in> perm \<Longrightarrow> MOST x. f x = x"
       
   101   unfolding perm_def MOST_iff_cofinite by simp
       
   102 
       
   103 lemma perm_id: "id \<in> perm"
       
   104   unfolding perm_def by simp
       
   105 
       
   106 lemma perm_comp:
       
   107   assumes f: "f \<in> perm" and g: "g \<in> perm"
       
   108   shows "(f \<circ> g) \<in> perm"
       
   109 apply (rule permI)
       
   110 apply (rule bij_comp)
       
   111 apply (rule perm_is_bij [OF g])
       
   112 apply (rule perm_is_bij [OF f])
       
   113 apply (rule MOST_rev_mp [OF perm_MOST [OF g]])
       
   114 apply (rule MOST_rev_mp [OF perm_MOST [OF f]])
       
   115 apply (simp)
       
   116 apply (simp add: perm_is_sort_respecting [OF f])
       
   117 apply (simp add: perm_is_sort_respecting [OF g])
       
   118 done
       
   119 
       
   120 lemma perm_inv:
       
   121   assumes f: "f \<in> perm"
       
   122   shows "(inv f) \<in> perm"
       
   123 apply (rule permI)
       
   124 apply (rule bij_imp_bij_inv)
       
   125 apply (rule perm_is_bij [OF f])
       
   126 apply (rule MOST_mono [OF perm_MOST [OF f]])
       
   127 apply (erule subst, rule inv_f_f)
       
   128 apply (rule bij_is_inj [OF perm_is_bij [OF f]])
       
   129 apply (rule perm_is_sort_respecting [OF f, THEN sym, THEN trans])
       
   130 apply (simp add: surj_f_inv_f [OF bij_is_surj [OF perm_is_bij [OF f]]])
       
   131 done
       
   132 
       
   133 lemma bij_Rep_perm: "bij (Rep_perm p)"
       
   134   using Rep_perm [of p] unfolding perm_def by simp
       
   135 
       
   136 lemma finite_Rep_perm: "finite {a. Rep_perm p a \<noteq> a}"
       
   137   using Rep_perm [of p] unfolding perm_def by simp
       
   138 
       
   139 lemma sort_of_Rep_perm: "sort_of (Rep_perm p a) = sort_of a"
       
   140   using Rep_perm [of p] unfolding perm_def by simp
       
   141 
       
   142 lemma Rep_perm_ext:
       
   143   "Rep_perm p1 = Rep_perm p2 \<Longrightarrow> p1 = p2"
       
   144   by (simp add: fun_eq_iff Rep_perm_inject [symmetric])
       
   145 
       
   146 instance perm :: size ..
       
   147 
       
   148 subsection {* Permutations form a group *}
       
   149 
       
   150 instantiation perm :: group_add
       
   151 begin
       
   152 
       
   153 definition
       
   154   "0 = Abs_perm id"
       
   155 
       
   156 definition
       
   157   "- p = Abs_perm (inv (Rep_perm p))"
       
   158 
       
   159 definition
       
   160   "p + q = Abs_perm (Rep_perm p \<circ> Rep_perm q)"
       
   161 
       
   162 definition
       
   163   "(p1::perm) - p2 = p1 + - p2"
       
   164 
       
   165 lemma Rep_perm_0: "Rep_perm 0 = id"
       
   166   unfolding zero_perm_def
       
   167   by (simp add: Abs_perm_inverse perm_id)
       
   168 
       
   169 lemma Rep_perm_add:
       
   170   "Rep_perm (p1 + p2) = Rep_perm p1 \<circ> Rep_perm p2"
       
   171   unfolding plus_perm_def
       
   172   by (simp add: Abs_perm_inverse perm_comp Rep_perm)
       
   173 
       
   174 lemma Rep_perm_uminus:
       
   175   "Rep_perm (- p) = inv (Rep_perm p)"
       
   176   unfolding uminus_perm_def
       
   177   by (simp add: Abs_perm_inverse perm_inv Rep_perm)
       
   178 
       
   179 instance
       
   180 apply default
       
   181 unfolding Rep_perm_inject [symmetric]
       
   182 unfolding minus_perm_def
       
   183 unfolding Rep_perm_add
       
   184 unfolding Rep_perm_uminus
       
   185 unfolding Rep_perm_0
       
   186 by (simp_all add: o_assoc inv_o_cancel [OF bij_is_inj [OF bij_Rep_perm]])
       
   187 
       
   188 end
       
   189 
       
   190 
       
   191 section {* Implementation of swappings *}
       
   192 
       
   193 definition
       
   194   swap :: "atom \<Rightarrow> atom \<Rightarrow> perm" ("'(_ \<rightleftharpoons> _')")
       
   195 where
       
   196   "(a \<rightleftharpoons> b) =
       
   197     Abs_perm (if sort_of a = sort_of b 
       
   198               then (\<lambda>c. if a = c then b else if b = c then a else c) 
       
   199               else id)"
       
   200 
       
   201 lemma Rep_perm_swap:
       
   202   "Rep_perm (a \<rightleftharpoons> b) =
       
   203     (if sort_of a = sort_of b 
       
   204      then (\<lambda>c. if a = c then b else if b = c then a else c)
       
   205      else id)"
       
   206 unfolding swap_def
       
   207 apply (rule Abs_perm_inverse)
       
   208 apply (rule permI)
       
   209 apply (auto simp add: bij_def inj_on_def surj_def)[1]
       
   210 apply (rule MOST_rev_mp [OF MOST_neq(1) [of a]])
       
   211 apply (rule MOST_rev_mp [OF MOST_neq(1) [of b]])
       
   212 apply (simp)
       
   213 apply (simp)
       
   214 done
       
   215 
       
   216 lemmas Rep_perm_simps =
       
   217   Rep_perm_0
       
   218   Rep_perm_add
       
   219   Rep_perm_uminus
       
   220   Rep_perm_swap
       
   221 
       
   222 lemma swap_different_sorts [simp]:
       
   223   "sort_of a \<noteq> sort_of b \<Longrightarrow> (a \<rightleftharpoons> b) = 0"
       
   224   by (rule Rep_perm_ext) (simp add: Rep_perm_simps)
       
   225 
       
   226 lemma swap_cancel:
       
   227   "(a \<rightleftharpoons> b) + (a \<rightleftharpoons> b) = 0"
       
   228   by (rule Rep_perm_ext) 
       
   229      (simp add: Rep_perm_simps fun_eq_iff)
       
   230 
       
   231 lemma swap_self [simp]:
       
   232   "(a \<rightleftharpoons> a) = 0"
       
   233   by (rule Rep_perm_ext, simp add: Rep_perm_simps fun_eq_iff)
       
   234 
       
   235 lemma minus_swap [simp]:
       
   236   "- (a \<rightleftharpoons> b) = (a \<rightleftharpoons> b)"
       
   237   by (rule minus_unique [OF swap_cancel])
       
   238 
       
   239 lemma swap_commute:
       
   240   "(a \<rightleftharpoons> b) = (b \<rightleftharpoons> a)"
       
   241   by (rule Rep_perm_ext)
       
   242      (simp add: Rep_perm_swap fun_eq_iff)
       
   243 
       
   244 lemma swap_triple:
       
   245   assumes "a \<noteq> b" and "c \<noteq> b"
       
   246   assumes "sort_of a = sort_of b" "sort_of b = sort_of c"
       
   247   shows "(a \<rightleftharpoons> c) + (b \<rightleftharpoons> c) + (a \<rightleftharpoons> c) = (a \<rightleftharpoons> b)"
       
   248   using assms
       
   249   by (rule_tac Rep_perm_ext)
       
   250      (auto simp add: Rep_perm_simps fun_eq_iff)
       
   251 
       
   252 
       
   253 section {* Permutation Types *}
       
   254 
       
   255 text {*
       
   256   Infix syntax for @{text permute} has higher precedence than
       
   257   addition, but lower than unary minus.
       
   258 *}
       
   259 
       
   260 class pt =
       
   261   fixes permute :: "perm \<Rightarrow> 'a \<Rightarrow> 'a" ("_ \<bullet> _" [76, 75] 75)
       
   262   assumes permute_zero [simp]: "0 \<bullet> x = x"
       
   263   assumes permute_plus [simp]: "(p + q) \<bullet> x = p \<bullet> (q \<bullet> x)"
       
   264 begin
       
   265 
       
   266 lemma permute_diff [simp]:
       
   267   shows "(p - q) \<bullet> x = p \<bullet> - q \<bullet> x"
       
   268   unfolding diff_minus by simp
       
   269 
       
   270 lemma permute_minus_cancel [simp]:
       
   271   shows "p \<bullet> - p \<bullet> x = x"
       
   272   and   "- p \<bullet> p \<bullet> x = x"
       
   273   unfolding permute_plus [symmetric] by simp_all
       
   274 
       
   275 lemma permute_swap_cancel [simp]:
       
   276   shows "(a \<rightleftharpoons> b) \<bullet> (a \<rightleftharpoons> b) \<bullet> x = x"
       
   277   unfolding permute_plus [symmetric]
       
   278   by (simp add: swap_cancel)
       
   279 
       
   280 lemma permute_swap_cancel2 [simp]:
       
   281   shows "(a \<rightleftharpoons> b) \<bullet> (b \<rightleftharpoons> a) \<bullet> x = x"
       
   282   unfolding permute_plus [symmetric]
       
   283   by (simp add: swap_commute)
       
   284 
       
   285 lemma inj_permute [simp]: 
       
   286   shows "inj (permute p)"
       
   287   by (rule inj_on_inverseI)
       
   288      (rule permute_minus_cancel)
       
   289 
       
   290 lemma surj_permute [simp]: 
       
   291   shows "surj (permute p)"
       
   292   by (rule surjI, rule permute_minus_cancel)
       
   293 
       
   294 lemma bij_permute [simp]: 
       
   295   shows "bij (permute p)"
       
   296   by (rule bijI [OF inj_permute surj_permute])
       
   297 
       
   298 lemma inv_permute: 
       
   299   shows "inv (permute p) = permute (- p)"
       
   300   by (rule inv_equality) (simp_all)
       
   301 
       
   302 lemma permute_minus: 
       
   303   shows "permute (- p) = inv (permute p)"
       
   304   by (simp add: inv_permute)
       
   305 
       
   306 lemma permute_eq_iff [simp]: 
       
   307   shows "p \<bullet> x = p \<bullet> y \<longleftrightarrow> x = y"
       
   308   by (rule inj_permute [THEN inj_eq])
       
   309 
       
   310 end
       
   311 
       
   312 subsection {* Permutations for atoms *}
       
   313 
       
   314 instantiation atom :: pt
       
   315 begin
       
   316 
       
   317 definition
       
   318   "p \<bullet> a = (Rep_perm p) a"
       
   319 
       
   320 instance 
       
   321 apply(default)
       
   322 apply(simp_all add: permute_atom_def Rep_perm_simps)
       
   323 done
       
   324 
       
   325 end
       
   326 
       
   327 lemma sort_of_permute [simp]:
       
   328   shows "sort_of (p \<bullet> a) = sort_of a"
       
   329   unfolding permute_atom_def by (rule sort_of_Rep_perm)
       
   330 
       
   331 lemma swap_atom:
       
   332   shows "(a \<rightleftharpoons> b) \<bullet> c =
       
   333            (if sort_of a = sort_of b
       
   334             then (if c = a then b else if c = b then a else c) else c)"
       
   335   unfolding permute_atom_def
       
   336   by (simp add: Rep_perm_swap)
       
   337 
       
   338 lemma swap_atom_simps [simp]:
       
   339   "sort_of a = sort_of b \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> a = b"
       
   340   "sort_of a = sort_of b \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> b = a"
       
   341   "c \<noteq> a \<Longrightarrow> c \<noteq> b \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> c = c"
       
   342   unfolding swap_atom by simp_all
       
   343 
       
   344 lemma expand_perm_eq:
       
   345   fixes p q :: "perm"
       
   346   shows "p = q \<longleftrightarrow> (\<forall>a::atom. p \<bullet> a = q \<bullet> a)"
       
   347   unfolding permute_atom_def
       
   348   by (metis Rep_perm_ext ext)
       
   349 
       
   350 
       
   351 subsection {* Permutations for permutations *}
       
   352 
       
   353 instantiation perm :: pt
       
   354 begin
       
   355 
       
   356 definition
       
   357   "p \<bullet> q = p + q - p"
       
   358 
       
   359 instance
       
   360 apply default
       
   361 apply (simp add: permute_perm_def)
       
   362 apply (simp add: permute_perm_def diff_minus minus_add add_assoc)
       
   363 done
       
   364 
       
   365 end
       
   366 
       
   367 lemma permute_self: 
       
   368   shows "p \<bullet> p = p"
       
   369   unfolding permute_perm_def 
       
   370   by (simp add: diff_minus add_assoc)
       
   371 
       
   372 lemma permute_eqvt:
       
   373   shows "p \<bullet> (q \<bullet> x) = (p \<bullet> q) \<bullet> (p \<bullet> x)"
       
   374   unfolding permute_perm_def by simp
       
   375 
       
   376 lemma zero_perm_eqvt:
       
   377   shows "p \<bullet> (0::perm) = 0"
       
   378   unfolding permute_perm_def by simp
       
   379 
       
   380 lemma add_perm_eqvt:
       
   381   fixes p p1 p2 :: perm
       
   382   shows "p \<bullet> (p1 + p2) = p \<bullet> p1 + p \<bullet> p2"
       
   383   unfolding permute_perm_def
       
   384   by (simp add: expand_perm_eq)
       
   385 
       
   386 lemma swap_eqvt:
       
   387   shows "p \<bullet> (a \<rightleftharpoons> b) = (p \<bullet> a \<rightleftharpoons> p \<bullet> b)"
       
   388   unfolding permute_perm_def
       
   389   by (auto simp add: swap_atom expand_perm_eq)
       
   390 
       
   391 lemma uminus_eqvt:
       
   392   fixes p q::"perm"
       
   393   shows "p \<bullet> (- q) = - (p \<bullet> q)"
       
   394   unfolding permute_perm_def
       
   395   by (simp add: diff_minus minus_add add_assoc)
       
   396 
       
   397 subsection {* Permutations for functions *}
       
   398 
       
   399 instantiation "fun" :: (pt, pt) pt
       
   400 begin
       
   401 
       
   402 definition
       
   403   "p \<bullet> f = (\<lambda>x. p \<bullet> (f (- p \<bullet> x)))"
       
   404 
       
   405 instance
       
   406 apply default
       
   407 apply (simp add: permute_fun_def)
       
   408 apply (simp add: permute_fun_def minus_add)
       
   409 done
       
   410 
       
   411 end
       
   412 
       
   413 lemma permute_fun_app_eq:
       
   414   shows "p \<bullet> (f x) = (p \<bullet> f) (p \<bullet> x)"
       
   415   unfolding permute_fun_def by simp
       
   416 
       
   417 
       
   418 subsection {* Permutations for booleans *}
       
   419 
       
   420 instantiation bool :: pt
       
   421 begin
       
   422 
       
   423 definition "p \<bullet> (b::bool) = b"
       
   424 
       
   425 instance
       
   426 apply(default) 
       
   427 apply(simp_all add: permute_bool_def)
       
   428 done
       
   429 
       
   430 end
       
   431 
       
   432 lemma Not_eqvt:
       
   433   shows "p \<bullet> (\<not> A) = (\<not> (p \<bullet> A))"
       
   434 by (simp add: permute_bool_def)
       
   435 
       
   436 lemma conj_eqvt:
       
   437   shows "p \<bullet> (A \<and> B) = ((p \<bullet> A) \<and> (p \<bullet> B))"
       
   438   by (simp add: permute_bool_def)
       
   439 
       
   440 lemma imp_eqvt:
       
   441   shows "p \<bullet> (A \<longrightarrow> B) = ((p \<bullet> A) \<longrightarrow> (p \<bullet> B))"
       
   442   by (simp add: permute_bool_def)
       
   443 
       
   444 lemma ex_eqvt:
       
   445   shows "p \<bullet> (\<exists>x. P x) = (\<exists>x. (p \<bullet> P) x)"
       
   446   unfolding permute_fun_def permute_bool_def
       
   447   by (auto, rule_tac x="p \<bullet> x" in exI, simp)
       
   448 
       
   449 lemma all_eqvt:
       
   450   shows "p \<bullet> (\<forall>x. P x) = (\<forall>x. (p \<bullet> P) x)"
       
   451   unfolding permute_fun_def permute_bool_def
       
   452   by (auto, drule_tac x="p \<bullet> x" in spec, simp)
       
   453 
       
   454 lemma permute_boolE:
       
   455   fixes P::"bool"
       
   456   shows "p \<bullet> P \<Longrightarrow> P"
       
   457   by (simp add: permute_bool_def)
       
   458 
       
   459 lemma permute_boolI:
       
   460   fixes P::"bool"
       
   461   shows "P \<Longrightarrow> p \<bullet> P"
       
   462   by(simp add: permute_bool_def)
       
   463 
       
   464 subsection {* Permutations for sets *}
       
   465 
       
   466 lemma permute_set_eq:
       
   467   fixes x::"'a::pt"
       
   468   and   p::"perm"
       
   469   shows "(p \<bullet> X) = {p \<bullet> x | x. x \<in> X}"
       
   470   unfolding permute_fun_def
       
   471   unfolding permute_bool_def
       
   472   apply(auto simp add: mem_def)
       
   473   apply(rule_tac x="- p \<bullet> x" in exI)
       
   474   apply(simp)
       
   475   done
       
   476 
       
   477 lemma permute_set_eq_image:
       
   478   shows "p \<bullet> X = permute p ` X"
       
   479   unfolding permute_set_eq by auto
       
   480 
       
   481 lemma permute_set_eq_vimage:
       
   482   shows "p \<bullet> X = permute (- p) -` X"
       
   483   unfolding permute_fun_def permute_bool_def
       
   484   unfolding vimage_def Collect_def mem_def ..
       
   485 
       
   486 lemma swap_set_not_in:
       
   487   assumes a: "a \<notin> S" "b \<notin> S"
       
   488   shows "(a \<rightleftharpoons> b) \<bullet> S = S"
       
   489   unfolding permute_set_eq
       
   490   using a by (auto simp add: swap_atom)
       
   491 
       
   492 lemma swap_set_in:
       
   493   assumes a: "a \<in> S" "b \<notin> S" "sort_of a = sort_of b"
       
   494   shows "(a \<rightleftharpoons> b) \<bullet> S \<noteq> S"
       
   495   unfolding permute_set_eq
       
   496   using a by (auto simp add: swap_atom)
       
   497 
       
   498 lemma mem_permute_iff:
       
   499   shows "(p \<bullet> x) \<in> (p \<bullet> X) \<longleftrightarrow> x \<in> X"
       
   500   unfolding mem_def permute_fun_def permute_bool_def
       
   501   by simp
       
   502 
       
   503 lemma mem_eqvt:
       
   504   shows "p \<bullet> (x \<in> A) \<longleftrightarrow> (p \<bullet> x) \<in> (p \<bullet> A)"
       
   505   unfolding mem_def
       
   506   by (simp add: permute_fun_app_eq)
       
   507 
       
   508 lemma empty_eqvt:
       
   509   shows "p \<bullet> {} = {}"
       
   510   unfolding empty_def Collect_def
       
   511   by (simp add: permute_fun_def permute_bool_def)
       
   512 
       
   513 lemma insert_eqvt:
       
   514   shows "p \<bullet> (insert x A) = insert (p \<bullet> x) (p \<bullet> A)"
       
   515   unfolding permute_set_eq_image image_insert ..
       
   516 
       
   517 
       
   518 subsection {* Permutations for units *}
       
   519 
       
   520 instantiation unit :: pt
       
   521 begin
       
   522 
       
   523 definition "p \<bullet> (u::unit) = u"
       
   524 
       
   525 instance 
       
   526 by (default) (simp_all add: permute_unit_def)
       
   527 
       
   528 end
       
   529 
       
   530 
       
   531 subsection {* Permutations for products *}
       
   532 
       
   533 instantiation prod :: (pt, pt) pt
       
   534 begin
       
   535 
       
   536 primrec 
       
   537   permute_prod 
       
   538 where
       
   539   Pair_eqvt: "p \<bullet> (x, y) = (p \<bullet> x, p \<bullet> y)"
       
   540 
       
   541 instance
       
   542 by default auto
       
   543 
       
   544 end
       
   545 
       
   546 subsection {* Permutations for sums *}
       
   547 
       
   548 instantiation sum :: (pt, pt) pt
       
   549 begin
       
   550 
       
   551 primrec 
       
   552   permute_sum 
       
   553 where
       
   554   "p \<bullet> (Inl x) = Inl (p \<bullet> x)"
       
   555 | "p \<bullet> (Inr y) = Inr (p \<bullet> y)"
       
   556 
       
   557 instance 
       
   558 by (default) (case_tac [!] x, simp_all)
       
   559 
       
   560 end
       
   561 
       
   562 subsection {* Permutations for lists *}
       
   563 
       
   564 instantiation list :: (pt) pt
       
   565 begin
       
   566 
       
   567 primrec 
       
   568   permute_list 
       
   569 where
       
   570   "p \<bullet> [] = []"
       
   571 | "p \<bullet> (x # xs) = p \<bullet> x # p \<bullet> xs"
       
   572 
       
   573 instance 
       
   574 by (default) (induct_tac [!] x, simp_all)
       
   575 
       
   576 end
       
   577 
       
   578 lemma set_eqvt:
       
   579   shows "p \<bullet> (set xs) = set (p \<bullet> xs)"
       
   580   by (induct xs) (simp_all add: empty_eqvt insert_eqvt)
       
   581 
       
   582 subsection {* Permutations for options *}
       
   583 
       
   584 instantiation option :: (pt) pt
       
   585 begin
       
   586 
       
   587 primrec 
       
   588   permute_option 
       
   589 where
       
   590   "p \<bullet> None = None"
       
   591 | "p \<bullet> (Some x) = Some (p \<bullet> x)"
       
   592 
       
   593 instance 
       
   594 by (default) (induct_tac [!] x, simp_all)
       
   595 
       
   596 end
       
   597 
       
   598 
       
   599 subsection {* Permutations for fsets *}
       
   600 
       
   601 lemma permute_fset_rsp[quot_respect]:
       
   602   shows "(op = ===> list_eq ===> list_eq) permute permute"
       
   603   unfolding fun_rel_def
       
   604   by (simp add: set_eqvt[symmetric])
       
   605 
       
   606 instantiation fset :: (pt) pt
       
   607 begin
       
   608 
       
   609 quotient_definition
       
   610   "permute_fset :: perm \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
       
   611 is
       
   612   "permute :: perm \<Rightarrow> 'a list \<Rightarrow> 'a list"
       
   613 
       
   614 instance 
       
   615 proof
       
   616   fix x :: "'a fset" and p q :: "perm"
       
   617   show "0 \<bullet> x = x" by (descending) (simp)
       
   618   show "(p + q) \<bullet> x = p \<bullet> q \<bullet> x" by (descending) (simp)
       
   619 qed
       
   620 
       
   621 end
       
   622 
       
   623 lemma permute_fset [simp]:
       
   624   fixes S::"('a::pt) fset"
       
   625   shows "(p \<bullet> {||}) = ({||} ::('a::pt) fset)"
       
   626   and   "(p \<bullet> insert_fset x S) = insert_fset (p \<bullet> x) (p \<bullet> S)"
       
   627   by (lifting permute_list.simps)
       
   628 
       
   629 
       
   630 
       
   631 subsection {* Permutations for @{typ char}, @{typ nat}, and @{typ int} *}
       
   632 
       
   633 instantiation char :: pt
       
   634 begin
       
   635 
       
   636 definition "p \<bullet> (c::char) = c"
       
   637 
       
   638 instance 
       
   639 by (default) (simp_all add: permute_char_def)
       
   640 
       
   641 end
       
   642 
       
   643 instantiation nat :: pt
       
   644 begin
       
   645 
       
   646 definition "p \<bullet> (n::nat) = n"
       
   647 
       
   648 instance 
       
   649 by (default) (simp_all add: permute_nat_def)
       
   650 
       
   651 end
       
   652 
       
   653 instantiation int :: pt
       
   654 begin
       
   655 
       
   656 definition "p \<bullet> (i::int) = i"
       
   657 
       
   658 instance 
       
   659 by (default) (simp_all add: permute_int_def)
       
   660 
       
   661 end
       
   662 
       
   663 
       
   664 section {* Pure types *}
       
   665 
       
   666 text {* Pure types will have always empty support. *}
       
   667 
       
   668 class pure = pt +
       
   669   assumes permute_pure: "p \<bullet> x = x"
       
   670 
       
   671 text {* Types @{typ unit} and @{typ bool} are pure. *}
       
   672 
       
   673 instance unit :: pure
       
   674 proof qed (rule permute_unit_def)
       
   675 
       
   676 instance bool :: pure
       
   677 proof qed (rule permute_bool_def)
       
   678 
       
   679 text {* Other type constructors preserve purity. *}
       
   680 
       
   681 instance "fun" :: (pure, pure) pure
       
   682 by default (simp add: permute_fun_def permute_pure)
       
   683 
       
   684 instance prod :: (pure, pure) pure
       
   685 by default (induct_tac x, simp add: permute_pure)
       
   686 
       
   687 instance sum :: (pure, pure) pure
       
   688 by default (induct_tac x, simp_all add: permute_pure)
       
   689 
       
   690 instance list :: (pure) pure
       
   691 by default (induct_tac x, simp_all add: permute_pure)
       
   692 
       
   693 instance option :: (pure) pure
       
   694 by default (induct_tac x, simp_all add: permute_pure)
       
   695 
       
   696 
       
   697 subsection {* Types @{typ char}, @{typ nat}, and @{typ int} *}
       
   698 
       
   699 instance char :: pure
       
   700 proof qed (rule permute_char_def)
       
   701 
       
   702 instance nat :: pure
       
   703 proof qed (rule permute_nat_def)
       
   704 
       
   705 instance int :: pure
       
   706 proof qed (rule permute_int_def)
       
   707 
       
   708 
       
   709 subsection {* Supp, Freshness and Supports *}
       
   710 
       
   711 context pt
       
   712 begin
       
   713 
       
   714 definition
       
   715   supp :: "'a \<Rightarrow> atom set"
       
   716 where
       
   717   "supp x = {a. infinite {b. (a \<rightleftharpoons> b) \<bullet> x \<noteq> x}}"
       
   718 
       
   719 end
       
   720 
       
   721 definition
       
   722   fresh :: "atom \<Rightarrow> 'a::pt \<Rightarrow> bool" ("_ \<sharp> _" [55, 55] 55)
       
   723 where   
       
   724   "a \<sharp> x \<equiv> a \<notin> supp x"
       
   725 
       
   726 lemma supp_conv_fresh: 
       
   727   shows "supp x = {a. \<not> a \<sharp> x}"
       
   728   unfolding fresh_def by simp
       
   729 
       
   730 lemma swap_rel_trans:
       
   731   assumes "sort_of a = sort_of b"
       
   732   assumes "sort_of b = sort_of c"
       
   733   assumes "(a \<rightleftharpoons> c) \<bullet> x = x"
       
   734   assumes "(b \<rightleftharpoons> c) \<bullet> x = x"
       
   735   shows "(a \<rightleftharpoons> b) \<bullet> x = x"
       
   736 proof (cases)
       
   737   assume "a = b \<or> c = b"
       
   738   with assms show "(a \<rightleftharpoons> b) \<bullet> x = x" by auto
       
   739 next
       
   740   assume *: "\<not> (a = b \<or> c = b)"
       
   741   have "((a \<rightleftharpoons> c) + (b \<rightleftharpoons> c) + (a \<rightleftharpoons> c)) \<bullet> x = x"
       
   742     using assms by simp
       
   743   also have "(a \<rightleftharpoons> c) + (b \<rightleftharpoons> c) + (a \<rightleftharpoons> c) = (a \<rightleftharpoons> b)"
       
   744     using assms * by (simp add: swap_triple)
       
   745   finally show "(a \<rightleftharpoons> b) \<bullet> x = x" .
       
   746 qed
       
   747 
       
   748 lemma swap_fresh_fresh:
       
   749   assumes a: "a \<sharp> x" 
       
   750   and     b: "b \<sharp> x"
       
   751   shows "(a \<rightleftharpoons> b) \<bullet> x = x"
       
   752 proof (cases)
       
   753   assume asm: "sort_of a = sort_of b" 
       
   754   have "finite {c. (a \<rightleftharpoons> c) \<bullet> x \<noteq> x}" "finite {c. (b \<rightleftharpoons> c) \<bullet> x \<noteq> x}" 
       
   755     using a b unfolding fresh_def supp_def by simp_all
       
   756   then have "finite ({c. (a \<rightleftharpoons> c) \<bullet> x \<noteq> x} \<union> {c. (b \<rightleftharpoons> c) \<bullet> x \<noteq> x})" by simp
       
   757   then obtain c 
       
   758     where "(a \<rightleftharpoons> c) \<bullet> x = x" "(b \<rightleftharpoons> c) \<bullet> x = x" "sort_of c = sort_of b"
       
   759     by (rule obtain_atom) (auto)
       
   760   then show "(a \<rightleftharpoons> b) \<bullet> x = x" using asm by (rule_tac swap_rel_trans) (simp_all)
       
   761 next
       
   762   assume "sort_of a \<noteq> sort_of b"
       
   763   then show "(a \<rightleftharpoons> b) \<bullet> x = x" by simp
       
   764 qed
       
   765 
       
   766 
       
   767 subsection {* supp and fresh are equivariant *}
       
   768 
       
   769 lemma finite_Collect_bij:
       
   770   assumes a: "bij f"
       
   771   shows "finite {x. P (f x)} = finite {x. P x}"
       
   772 by (metis a finite_vimage_iff vimage_Collect_eq)
       
   773 
       
   774 lemma fresh_permute_iff:
       
   775   shows "(p \<bullet> a) \<sharp> (p \<bullet> x) \<longleftrightarrow> a \<sharp> x"
       
   776 proof -
       
   777   have "(p \<bullet> a) \<sharp> (p \<bullet> x) \<longleftrightarrow> finite {b. (p \<bullet> a \<rightleftharpoons> b) \<bullet> p \<bullet> x \<noteq> p \<bullet> x}"
       
   778     unfolding fresh_def supp_def by simp
       
   779   also have "\<dots> \<longleftrightarrow> finite {b. (p \<bullet> a \<rightleftharpoons> p \<bullet> b) \<bullet> p \<bullet> x \<noteq> p \<bullet> x}"
       
   780     using bij_permute by (rule finite_Collect_bij[symmetric])
       
   781   also have "\<dots> \<longleftrightarrow> finite {b. p \<bullet> (a \<rightleftharpoons> b) \<bullet> x \<noteq> p \<bullet> x}"
       
   782     by (simp only: permute_eqvt [of p] swap_eqvt)
       
   783   also have "\<dots> \<longleftrightarrow> finite {b. (a \<rightleftharpoons> b) \<bullet> x \<noteq> x}"
       
   784     by (simp only: permute_eq_iff)
       
   785   also have "\<dots> \<longleftrightarrow> a \<sharp> x"
       
   786     unfolding fresh_def supp_def by simp
       
   787   finally show "(p \<bullet> a) \<sharp> (p \<bullet> x) \<longleftrightarrow> a \<sharp> x" .
       
   788 qed
       
   789 
       
   790 lemma fresh_eqvt:
       
   791   shows "p \<bullet> (a \<sharp> x) = (p \<bullet> a) \<sharp> (p \<bullet> x)"
       
   792   unfolding permute_bool_def
       
   793   by (simp add: fresh_permute_iff)
       
   794 
       
   795 lemma supp_eqvt:
       
   796   fixes  p  :: "perm"
       
   797   and    x   :: "'a::pt"
       
   798   shows "p \<bullet> (supp x) = supp (p \<bullet> x)"
       
   799   unfolding supp_conv_fresh
       
   800   unfolding Collect_def
       
   801   unfolding permute_fun_def
       
   802   by (simp add: Not_eqvt fresh_eqvt)
       
   803 
       
   804 subsection {* supports *}
       
   805 
       
   806 definition
       
   807   supports :: "atom set \<Rightarrow> 'a::pt \<Rightarrow> bool" (infixl "supports" 80)
       
   808 where  
       
   809   "S supports x \<equiv> \<forall>a b. (a \<notin> S \<and> b \<notin> S \<longrightarrow> (a \<rightleftharpoons> b) \<bullet> x = x)"
       
   810 
       
   811 lemma supp_is_subset:
       
   812   fixes S :: "atom set"
       
   813   and   x :: "'a::pt"
       
   814   assumes a1: "S supports x"
       
   815   and     a2: "finite S"
       
   816   shows "(supp x) \<subseteq> S"
       
   817 proof (rule ccontr)
       
   818   assume "\<not> (supp x \<subseteq> S)"
       
   819   then obtain a where b1: "a \<in> supp x" and b2: "a \<notin> S" by auto
       
   820   from a1 b2 have "\<forall>b. b \<notin> S \<longrightarrow> (a \<rightleftharpoons> b) \<bullet> x = x" unfolding supports_def by auto
       
   821   then have "{b. (a \<rightleftharpoons> b) \<bullet> x \<noteq> x} \<subseteq> S" by auto
       
   822   with a2 have "finite {b. (a \<rightleftharpoons> b)\<bullet>x \<noteq> x}" by (simp add: finite_subset)
       
   823   then have "a \<notin> (supp x)" unfolding supp_def by simp
       
   824   with b1 show False by simp
       
   825 qed
       
   826 
       
   827 lemma supports_finite:
       
   828   fixes S :: "atom set"
       
   829   and   x :: "'a::pt"
       
   830   assumes a1: "S supports x"
       
   831   and     a2: "finite S"
       
   832   shows "finite (supp x)"
       
   833 proof -
       
   834   have "(supp x) \<subseteq> S" using a1 a2 by (rule supp_is_subset)
       
   835   then show "finite (supp x)" using a2 by (simp add: finite_subset)
       
   836 qed
       
   837 
       
   838 lemma supp_supports:
       
   839   fixes x :: "'a::pt"
       
   840   shows "(supp x) supports x"
       
   841 unfolding supports_def
       
   842 proof (intro strip)
       
   843   fix a b
       
   844   assume "a \<notin> (supp x) \<and> b \<notin> (supp x)"
       
   845   then have "a \<sharp> x" and "b \<sharp> x" by (simp_all add: fresh_def)
       
   846   then show "(a \<rightleftharpoons> b) \<bullet> x = x" by (simp add: swap_fresh_fresh)
       
   847 qed
       
   848 
       
   849 lemma supp_is_least_supports:
       
   850   fixes S :: "atom set"
       
   851   and   x :: "'a::pt"
       
   852   assumes  a1: "S supports x"
       
   853   and      a2: "finite S"
       
   854   and      a3: "\<And>S'. finite S' \<Longrightarrow> (S' supports x) \<Longrightarrow> S \<subseteq> S'"
       
   855   shows "(supp x) = S"
       
   856 proof (rule equalityI)
       
   857   show "(supp x) \<subseteq> S" using a1 a2 by (rule supp_is_subset)
       
   858   with a2 have fin: "finite (supp x)" by (rule rev_finite_subset)
       
   859   have "(supp x) supports x" by (rule supp_supports)
       
   860   with fin a3 show "S \<subseteq> supp x" by blast
       
   861 qed
       
   862 
       
   863 lemma subsetCI: 
       
   864   shows "(\<And>x. x \<in> A \<Longrightarrow> x \<notin> B \<Longrightarrow> False) \<Longrightarrow> A \<subseteq> B"
       
   865   by auto
       
   866 
       
   867 lemma finite_supp_unique:
       
   868   assumes a1: "S supports x"
       
   869   assumes a2: "finite S"
       
   870   assumes a3: "\<And>a b. \<lbrakk>a \<in> S; b \<notin> S; sort_of a = sort_of b\<rbrakk> \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> x \<noteq> x"
       
   871   shows "(supp x) = S"
       
   872   using a1 a2
       
   873 proof (rule supp_is_least_supports)
       
   874   fix S'
       
   875   assume "finite S'" and "S' supports x"
       
   876   show "S \<subseteq> S'"
       
   877   proof (rule subsetCI)
       
   878     fix a
       
   879     assume "a \<in> S" and "a \<notin> S'"
       
   880     have "finite (S \<union> S')"
       
   881       using `finite S` `finite S'` by simp
       
   882     then obtain b where "b \<notin> S \<union> S'" and "sort_of b = sort_of a"
       
   883       by (rule obtain_atom)
       
   884     then have "b \<notin> S" and "b \<notin> S'"  and "sort_of a = sort_of b"
       
   885       by simp_all
       
   886     then have "(a \<rightleftharpoons> b) \<bullet> x = x"
       
   887       using `a \<notin> S'` `S' supports x` by (simp add: supports_def)
       
   888     moreover have "(a \<rightleftharpoons> b) \<bullet> x \<noteq> x"
       
   889       using `a \<in> S` `b \<notin> S` `sort_of a = sort_of b`
       
   890       by (rule a3)
       
   891     ultimately show "False" by simp
       
   892   qed
       
   893 qed
       
   894 
       
   895 section {* Support w.r.t. relations *}
       
   896 
       
   897 text {* 
       
   898   This definition is used for unquotient types, where
       
   899   alpha-equivalence does not coincide with equality.
       
   900 *}
       
   901 
       
   902 definition 
       
   903   "supp_rel R x = {a. infinite {b. \<not>(R ((a \<rightleftharpoons> b) \<bullet> x) x)}}"
       
   904 
       
   905 
       
   906 
       
   907 section {* Finitely-supported types *}
       
   908 
       
   909 class fs = pt +
       
   910   assumes finite_supp: "finite (supp x)"
       
   911 
       
   912 lemma pure_supp: 
       
   913   shows "supp (x::'a::pure) = {}"
       
   914   unfolding supp_def by (simp add: permute_pure)
       
   915 
       
   916 lemma pure_fresh:
       
   917   fixes x::"'a::pure"
       
   918   shows "a \<sharp> x"
       
   919   unfolding fresh_def by (simp add: pure_supp)
       
   920 
       
   921 instance pure < fs
       
   922 by default (simp add: pure_supp)
       
   923 
       
   924 
       
   925 subsection  {* Type @{typ atom} is finitely-supported. *}
       
   926 
       
   927 lemma supp_atom:
       
   928   shows "supp a = {a}"
       
   929 apply (rule finite_supp_unique)
       
   930 apply (clarsimp simp add: supports_def)
       
   931 apply simp
       
   932 apply simp
       
   933 done
       
   934 
       
   935 lemma fresh_atom: 
       
   936   shows "a \<sharp> b \<longleftrightarrow> a \<noteq> b"
       
   937   unfolding fresh_def supp_atom by simp
       
   938 
       
   939 instance atom :: fs
       
   940 by default (simp add: supp_atom)
       
   941 
       
   942 section {* Support for finite sets of atoms *}
       
   943 
       
   944 
       
   945 lemma supp_finite_atom_set:
       
   946   fixes S::"atom set"
       
   947   assumes "finite S"
       
   948   shows "supp S = S"
       
   949   apply(rule finite_supp_unique)
       
   950   apply(simp add: supports_def)
       
   951   apply(simp add: swap_set_not_in)
       
   952   apply(rule assms)
       
   953   apply(simp add: swap_set_in)
       
   954 done
       
   955 
       
   956 section {* Type @{typ perm} is finitely-supported. *}
       
   957 
       
   958 lemma perm_swap_eq:
       
   959   shows "(a \<rightleftharpoons> b) \<bullet> p = p \<longleftrightarrow> (p \<bullet> (a \<rightleftharpoons> b)) = (a \<rightleftharpoons> b)"
       
   960 unfolding permute_perm_def
       
   961 by (metis add_diff_cancel minus_perm_def)
       
   962 
       
   963 lemma supports_perm: 
       
   964   shows "{a. p \<bullet> a \<noteq> a} supports p"
       
   965   unfolding supports_def
       
   966   unfolding perm_swap_eq
       
   967   by (simp add: swap_eqvt)
       
   968 
       
   969 lemma finite_perm_lemma: 
       
   970   shows "finite {a::atom. p \<bullet> a \<noteq> a}"
       
   971   using finite_Rep_perm [of p]
       
   972   unfolding permute_atom_def .
       
   973 
       
   974 lemma supp_perm:
       
   975   shows "supp p = {a. p \<bullet> a \<noteq> a}"
       
   976 apply (rule finite_supp_unique)
       
   977 apply (rule supports_perm)
       
   978 apply (rule finite_perm_lemma)
       
   979 apply (simp add: perm_swap_eq swap_eqvt)
       
   980 apply (auto simp add: expand_perm_eq swap_atom)
       
   981 done
       
   982 
       
   983 lemma fresh_perm: 
       
   984   shows "a \<sharp> p \<longleftrightarrow> p \<bullet> a = a"
       
   985   unfolding fresh_def 
       
   986   by (simp add: supp_perm)
       
   987 
       
   988 lemma supp_swap:
       
   989   shows "supp (a \<rightleftharpoons> b) = (if a = b \<or> sort_of a \<noteq> sort_of b then {} else {a, b})"
       
   990   by (auto simp add: supp_perm swap_atom)
       
   991 
       
   992 lemma fresh_zero_perm: 
       
   993   shows "a \<sharp> (0::perm)"
       
   994   unfolding fresh_perm by simp
       
   995 
       
   996 lemma supp_zero_perm: 
       
   997   shows "supp (0::perm) = {}"
       
   998   unfolding supp_perm by simp
       
   999 
       
  1000 lemma fresh_plus_perm:
       
  1001   fixes p q::perm
       
  1002   assumes "a \<sharp> p" "a \<sharp> q"
       
  1003   shows "a \<sharp> (p + q)"
       
  1004   using assms
       
  1005   unfolding fresh_def
       
  1006   by (auto simp add: supp_perm)
       
  1007 
       
  1008 lemma supp_plus_perm:
       
  1009   fixes p q::perm
       
  1010   shows "supp (p + q) \<subseteq> supp p \<union> supp q"
       
  1011   by (auto simp add: supp_perm)
       
  1012 
       
  1013 lemma fresh_minus_perm:
       
  1014   fixes p::perm
       
  1015   shows "a \<sharp> (- p) \<longleftrightarrow> a \<sharp> p"
       
  1016   unfolding fresh_def
       
  1017   unfolding supp_perm
       
  1018   apply(simp)
       
  1019   apply(metis permute_minus_cancel)
       
  1020   done
       
  1021 
       
  1022 lemma supp_minus_perm:
       
  1023   fixes p::perm
       
  1024   shows "supp (- p) = supp p"
       
  1025   unfolding supp_conv_fresh
       
  1026   by (simp add: fresh_minus_perm)
       
  1027 
       
  1028 instance perm :: fs
       
  1029 by default (simp add: supp_perm finite_perm_lemma)
       
  1030 
       
  1031 lemma plus_perm_eq:
       
  1032   fixes p q::"perm"
       
  1033   assumes asm: "supp p \<inter> supp q = {}"
       
  1034   shows "p + q  = q + p"
       
  1035 unfolding expand_perm_eq
       
  1036 proof
       
  1037   fix a::"atom"
       
  1038   show "(p + q) \<bullet> a = (q + p) \<bullet> a"
       
  1039   proof -
       
  1040     { assume "a \<notin> supp p" "a \<notin> supp q"
       
  1041       then have "(p + q) \<bullet> a = (q + p) \<bullet> a" 
       
  1042 	by (simp add: supp_perm)
       
  1043     }
       
  1044     moreover
       
  1045     { assume a: "a \<in> supp p" "a \<notin> supp q"
       
  1046       then have "p \<bullet> a \<in> supp p" by (simp add: supp_perm)
       
  1047       then have "p \<bullet> a \<notin> supp q" using asm by auto
       
  1048       with a have "(p + q) \<bullet> a = (q + p) \<bullet> a" 
       
  1049 	by (simp add: supp_perm)
       
  1050     }
       
  1051     moreover
       
  1052     { assume a: "a \<notin> supp p" "a \<in> supp q"
       
  1053       then have "q \<bullet> a \<in> supp q" by (simp add: supp_perm)
       
  1054       then have "q \<bullet> a \<notin> supp p" using asm by auto 
       
  1055       with a have "(p + q) \<bullet> a = (q + p) \<bullet> a" 
       
  1056 	by (simp add: supp_perm)
       
  1057     }
       
  1058     ultimately show "(p + q) \<bullet> a = (q + p) \<bullet> a" 
       
  1059       using asm by blast
       
  1060   qed
       
  1061 qed
       
  1062 
       
  1063 section {* Finite Support instances for other types *}
       
  1064 
       
  1065 
       
  1066 subsection {* Type @{typ "'a \<times> 'b"} is finitely-supported. *}
       
  1067 
       
  1068 lemma supp_Pair: 
       
  1069   shows "supp (x, y) = supp x \<union> supp y"
       
  1070   by (simp add: supp_def Collect_imp_eq Collect_neg_eq)
       
  1071 
       
  1072 lemma fresh_Pair: 
       
  1073   shows "a \<sharp> (x, y) \<longleftrightarrow> a \<sharp> x \<and> a \<sharp> y"
       
  1074   by (simp add: fresh_def supp_Pair)
       
  1075 
       
  1076 lemma supp_Unit:
       
  1077   shows "supp () = {}"
       
  1078   by (simp add: supp_def)
       
  1079 
       
  1080 lemma fresh_Unit:
       
  1081   shows "a \<sharp> ()"
       
  1082   by (simp add: fresh_def supp_Unit)
       
  1083 
       
  1084 instance prod :: (fs, fs) fs
       
  1085 apply default
       
  1086 apply (induct_tac x)
       
  1087 apply (simp add: supp_Pair finite_supp)
       
  1088 done
       
  1089 
       
  1090 
       
  1091 subsection {* Type @{typ "'a + 'b"} is finitely supported *}
       
  1092 
       
  1093 lemma supp_Inl: 
       
  1094   shows "supp (Inl x) = supp x"
       
  1095   by (simp add: supp_def)
       
  1096 
       
  1097 lemma supp_Inr: 
       
  1098   shows "supp (Inr x) = supp x"
       
  1099   by (simp add: supp_def)
       
  1100 
       
  1101 lemma fresh_Inl: 
       
  1102   shows "a \<sharp> Inl x \<longleftrightarrow> a \<sharp> x"
       
  1103   by (simp add: fresh_def supp_Inl)
       
  1104 
       
  1105 lemma fresh_Inr: 
       
  1106   shows "a \<sharp> Inr y \<longleftrightarrow> a \<sharp> y"
       
  1107   by (simp add: fresh_def supp_Inr)
       
  1108 
       
  1109 instance sum :: (fs, fs) fs
       
  1110 apply default
       
  1111 apply (induct_tac x)
       
  1112 apply (simp_all add: supp_Inl supp_Inr finite_supp)
       
  1113 done
       
  1114 
       
  1115 
       
  1116 subsection {* Type @{typ "'a option"} is finitely supported *}
       
  1117 
       
  1118 lemma supp_None: 
       
  1119   shows "supp None = {}"
       
  1120 by (simp add: supp_def)
       
  1121 
       
  1122 lemma supp_Some: 
       
  1123   shows "supp (Some x) = supp x"
       
  1124   by (simp add: supp_def)
       
  1125 
       
  1126 lemma fresh_None: 
       
  1127   shows "a \<sharp> None"
       
  1128   by (simp add: fresh_def supp_None)
       
  1129 
       
  1130 lemma fresh_Some: 
       
  1131   shows "a \<sharp> Some x \<longleftrightarrow> a \<sharp> x"
       
  1132   by (simp add: fresh_def supp_Some)
       
  1133 
       
  1134 instance option :: (fs) fs
       
  1135 apply default
       
  1136 apply (induct_tac x)
       
  1137 apply (simp_all add: supp_None supp_Some finite_supp)
       
  1138 done
       
  1139 
       
  1140 
       
  1141 subsubsection {* Type @{typ "'a list"} is finitely supported *}
       
  1142 
       
  1143 lemma supp_Nil: 
       
  1144   shows "supp [] = {}"
       
  1145   by (simp add: supp_def)
       
  1146 
       
  1147 lemma supp_Cons: 
       
  1148   shows "supp (x # xs) = supp x \<union> supp xs"
       
  1149 by (simp add: supp_def Collect_imp_eq Collect_neg_eq)
       
  1150 
       
  1151 lemma fresh_Nil: 
       
  1152   shows "a \<sharp> []"
       
  1153   by (simp add: fresh_def supp_Nil)
       
  1154 
       
  1155 lemma fresh_Cons: 
       
  1156   shows "a \<sharp> (x # xs) \<longleftrightarrow> a \<sharp> x \<and> a \<sharp> xs"
       
  1157   by (simp add: fresh_def supp_Cons)
       
  1158 
       
  1159 instance list :: (fs) fs
       
  1160 apply default
       
  1161 apply (induct_tac x)
       
  1162 apply (simp_all add: supp_Nil supp_Cons finite_supp)
       
  1163 done
       
  1164 
       
  1165 
       
  1166 section {* Support and Freshness for Applications *}
       
  1167 
       
  1168 lemma fresh_conv_MOST: 
       
  1169   shows "a \<sharp> x \<longleftrightarrow> (MOST b. (a \<rightleftharpoons> b) \<bullet> x = x)"
       
  1170   unfolding fresh_def supp_def 
       
  1171   unfolding MOST_iff_cofinite by simp
       
  1172 
       
  1173 lemma supp_subset_fresh:
       
  1174   assumes a: "\<And>a. a \<sharp> x \<Longrightarrow> a \<sharp> y"
       
  1175   shows "supp y \<subseteq> supp x"
       
  1176   using a
       
  1177   unfolding fresh_def
       
  1178   by blast
       
  1179 
       
  1180 lemma fresh_fun_app:
       
  1181   assumes "a \<sharp> f" and "a \<sharp> x" 
       
  1182   shows "a \<sharp> f x"
       
  1183   using assms
       
  1184   unfolding fresh_conv_MOST
       
  1185   unfolding permute_fun_app_eq 
       
  1186   by (elim MOST_rev_mp, simp)
       
  1187 
       
  1188 lemma supp_fun_app:
       
  1189   shows "supp (f x) \<subseteq> (supp f) \<union> (supp x)"
       
  1190   using fresh_fun_app
       
  1191   unfolding fresh_def
       
  1192   by auto
       
  1193 
       
  1194 text {* Support of Equivariant Functions *}
       
  1195 
       
  1196 lemma supp_fun_eqvt:
       
  1197   assumes a: "\<And>p. p \<bullet> f = f"
       
  1198   shows "supp f = {}"
       
  1199   unfolding supp_def 
       
  1200   using a by simp
       
  1201 
       
  1202 lemma fresh_fun_eqvt_app:
       
  1203   assumes a: "\<And>p. p \<bullet> f = f"
       
  1204   shows "a \<sharp> x \<Longrightarrow> a \<sharp> f x"
       
  1205 proof -
       
  1206   from a have "supp f = {}" by (simp add: supp_fun_eqvt)
       
  1207   then show "a \<sharp> x \<Longrightarrow> a \<sharp> f x"
       
  1208     unfolding fresh_def
       
  1209     using supp_fun_app by auto
       
  1210 qed
       
  1211 
       
  1212 
       
  1213 section {* Support of Finite Sets of Finitely Supported Elements *}
       
  1214 
       
  1215 lemma Union_fresh:
       
  1216   shows "a \<sharp> S \<Longrightarrow> a \<sharp> (\<Union>x \<in> S. supp x)"
       
  1217   unfolding Union_image_eq[symmetric]
       
  1218   apply(rule_tac f="\<lambda>S. \<Union> supp ` S" in fresh_fun_eqvt_app)
       
  1219   apply(simp add: permute_fun_def UNION_def Collect_def Bex_def ex_eqvt mem_def conj_eqvt)
       
  1220   apply(subst permute_fun_app_eq)
       
  1221   back
       
  1222   apply(simp add: supp_eqvt)
       
  1223   apply(assumption)
       
  1224   done
       
  1225 
       
  1226 lemma Union_supports_set:
       
  1227   shows "(\<Union>x \<in> S. supp x) supports S"
       
  1228 proof -
       
  1229   { fix a b
       
  1230     have "\<forall>x \<in> S. (a \<rightleftharpoons> b) \<bullet> x = x \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> S = S"
       
  1231       unfolding permute_set_eq by force
       
  1232   }
       
  1233   then show "(\<Union>x \<in> S. supp x) supports S"
       
  1234     unfolding supports_def 
       
  1235     by (simp add: fresh_def[symmetric] swap_fresh_fresh)
       
  1236 qed
       
  1237 
       
  1238 lemma Union_of_finite_supp_sets:
       
  1239   fixes S::"('a::fs set)"
       
  1240   assumes fin: "finite S"   
       
  1241   shows "finite (\<Union>x\<in>S. supp x)"
       
  1242   using fin by (induct) (auto simp add: finite_supp)
       
  1243 
       
  1244 lemma Union_included_in_supp:
       
  1245   fixes S::"('a::fs set)"
       
  1246   assumes fin: "finite S"
       
  1247   shows "(\<Union>x\<in>S. supp x) \<subseteq> supp S"
       
  1248 proof -
       
  1249   have "(\<Union>x\<in>S. supp x) = supp (\<Union>x\<in>S. supp x)"
       
  1250     by (rule supp_finite_atom_set[symmetric])
       
  1251        (rule Union_of_finite_supp_sets[OF fin])
       
  1252   also have "\<dots> \<subseteq> supp S"
       
  1253     by (rule supp_subset_fresh)
       
  1254        (simp add: Union_fresh)
       
  1255   finally show "(\<Union>x\<in>S. supp x) \<subseteq> supp S" .
       
  1256 qed
       
  1257 
       
  1258 lemma supp_of_finite_sets:
       
  1259   fixes S::"('a::fs set)"
       
  1260   assumes fin: "finite S"
       
  1261   shows "(supp S) = (\<Union>x\<in>S. supp x)"
       
  1262 apply(rule subset_antisym)
       
  1263 apply(rule supp_is_subset)
       
  1264 apply(rule Union_supports_set)
       
  1265 apply(rule Union_of_finite_supp_sets[OF fin])
       
  1266 apply(rule Union_included_in_supp[OF fin])
       
  1267 done
       
  1268 
       
  1269 lemma finite_sets_supp:
       
  1270   fixes S::"('a::fs set)"
       
  1271   assumes "finite S"
       
  1272   shows "finite (supp S)"
       
  1273 using assms
       
  1274 by (simp only: supp_of_finite_sets Union_of_finite_supp_sets)
       
  1275 
       
  1276 lemma supp_of_finite_union:
       
  1277   fixes S T::"('a::fs) set"
       
  1278   assumes fin1: "finite S"
       
  1279   and     fin2: "finite T"
       
  1280   shows "supp (S \<union> T) = supp S \<union> supp T"
       
  1281   using fin1 fin2
       
  1282   by (simp add: supp_of_finite_sets)
       
  1283 
       
  1284 lemma supp_of_finite_insert:
       
  1285   fixes S::"('a::fs) set"
       
  1286   assumes fin:  "finite S"
       
  1287   shows "supp (insert x S) = supp x \<union> supp S"
       
  1288   using fin
       
  1289   by (simp add: supp_of_finite_sets)
       
  1290 
       
  1291 
       
  1292 subsection {* Type @{typ "'a fset"} is finitely supported *}
       
  1293 
       
  1294 lemma fset_eqvt: 
       
  1295   shows "p \<bullet> (fset S) = fset (p \<bullet> S)"
       
  1296   by (lifting set_eqvt)
       
  1297 
       
  1298 lemma supp_fset [simp]:
       
  1299   shows "supp (fset S) = supp S"
       
  1300   unfolding supp_def
       
  1301   by (simp add: fset_eqvt fset_cong)
       
  1302 
       
  1303 lemma supp_empty_fset [simp]:
       
  1304   shows "supp {||} = {}"
       
  1305   unfolding supp_def
       
  1306   by simp
       
  1307 
       
  1308 lemma supp_insert_fset [simp]:
       
  1309   fixes x::"'a::fs"
       
  1310   and   S::"'a fset"
       
  1311   shows "supp (insert_fset x S) = supp x \<union> supp S"
       
  1312   apply(subst supp_fset[symmetric])
       
  1313   apply(simp add: supp_fset supp_of_finite_insert)
       
  1314   done
       
  1315 
       
  1316 lemma fset_finite_supp:
       
  1317   fixes S::"('a::fs) fset"
       
  1318   shows "finite (supp S)"
       
  1319   by (induct S) (simp_all add: finite_supp)
       
  1320 
       
  1321 
       
  1322 instance fset :: (fs) fs
       
  1323   apply (default)
       
  1324   apply (rule fset_finite_supp)
       
  1325   done
       
  1326 
       
  1327 
       
  1328 section {* Fresh-Star *}
       
  1329 
       
  1330 
       
  1331 text {* The fresh-star generalisation of fresh is used in strong
       
  1332   induction principles. *}
       
  1333 
       
  1334 definition 
       
  1335   fresh_star :: "atom set \<Rightarrow> 'a::pt \<Rightarrow> bool" ("_ \<sharp>* _" [80,80] 80)
       
  1336 where 
       
  1337   "as \<sharp>* x \<equiv> \<forall>a \<in> as. a \<sharp> x"
       
  1338 
       
  1339 lemma fresh_star_supp_conv:
       
  1340   shows "supp x \<sharp>* y \<Longrightarrow> supp y \<sharp>* x"
       
  1341 by (auto simp add: fresh_star_def fresh_def)
       
  1342 
       
  1343 lemma fresh_star_prod:
       
  1344   fixes as::"atom set"
       
  1345   shows "as \<sharp>* (x, y) = (as \<sharp>* x \<and> as \<sharp>* y)" 
       
  1346   by (auto simp add: fresh_star_def fresh_Pair)
       
  1347 
       
  1348 lemma fresh_star_union:
       
  1349   shows "(as \<union> bs) \<sharp>* x = (as \<sharp>* x \<and> bs \<sharp>* x)"
       
  1350   by (auto simp add: fresh_star_def)
       
  1351 
       
  1352 lemma fresh_star_insert:
       
  1353   shows "(insert a as) \<sharp>* x = (a \<sharp> x \<and> as \<sharp>* x)"
       
  1354   by (auto simp add: fresh_star_def)
       
  1355 
       
  1356 lemma fresh_star_Un_elim:
       
  1357   "((as \<union> bs) \<sharp>* x \<Longrightarrow> PROP C) \<equiv> (as \<sharp>* x \<Longrightarrow> bs \<sharp>* x \<Longrightarrow> PROP C)"
       
  1358   unfolding fresh_star_def
       
  1359   apply(rule)
       
  1360   apply(erule meta_mp)
       
  1361   apply(auto)
       
  1362   done
       
  1363 
       
  1364 lemma fresh_star_insert_elim:
       
  1365   "(insert a as \<sharp>* x \<Longrightarrow> PROP C) \<equiv> (a \<sharp> x \<Longrightarrow> as \<sharp>* x \<Longrightarrow> PROP C)"
       
  1366   unfolding fresh_star_def
       
  1367   by rule (simp_all add: fresh_star_def)
       
  1368 
       
  1369 lemma fresh_star_empty_elim:
       
  1370   "({} \<sharp>* x \<Longrightarrow> PROP C) \<equiv> PROP C"
       
  1371   by (simp add: fresh_star_def)
       
  1372 
       
  1373 lemma fresh_star_unit_elim: 
       
  1374   shows "(a \<sharp>* () \<Longrightarrow> PROP C) \<equiv> PROP C"
       
  1375   by (simp add: fresh_star_def fresh_Unit) 
       
  1376 
       
  1377 lemma fresh_star_prod_elim: 
       
  1378   shows "(a \<sharp>* (x, y) \<Longrightarrow> PROP C) \<equiv> (a \<sharp>* x \<Longrightarrow> a \<sharp>* y \<Longrightarrow> PROP C)"
       
  1379   by (rule, simp_all add: fresh_star_prod)
       
  1380 
       
  1381 lemma fresh_star_zero:
       
  1382   shows "as \<sharp>* (0::perm)"
       
  1383   unfolding fresh_star_def
       
  1384   by (simp add: fresh_zero_perm)
       
  1385 
       
  1386 lemma fresh_star_plus:
       
  1387   fixes p q::perm
       
  1388   shows "\<lbrakk>a \<sharp>* p;  a \<sharp>* q\<rbrakk> \<Longrightarrow> a \<sharp>* (p + q)"
       
  1389   unfolding fresh_star_def
       
  1390   by (simp add: fresh_plus_perm)
       
  1391 
       
  1392 lemma fresh_star_permute_iff:
       
  1393   shows "(p \<bullet> a) \<sharp>* (p \<bullet> x) \<longleftrightarrow> a \<sharp>* x"
       
  1394   unfolding fresh_star_def
       
  1395   by (metis mem_permute_iff permute_minus_cancel(1) fresh_permute_iff)
       
  1396 
       
  1397 lemma fresh_star_eqvt:
       
  1398   shows "(p \<bullet> (as \<sharp>* x)) = (p \<bullet> as) \<sharp>* (p \<bullet> x)"
       
  1399 unfolding fresh_star_def
       
  1400 unfolding Ball_def
       
  1401 apply(simp add: all_eqvt)
       
  1402 apply(subst permute_fun_def)
       
  1403 apply(simp add: imp_eqvt fresh_eqvt mem_eqvt)
       
  1404 done
       
  1405 
       
  1406 
       
  1407 section {* Induction principle for permutations *}
       
  1408 
       
  1409 
       
  1410 lemma perm_struct_induct[consumes 1, case_names zero swap]:
       
  1411   assumes S: "supp p \<subseteq> S"
       
  1412   and zero: "P 0"
       
  1413   and swap: "\<And>p a b. \<lbrakk>P p; supp p \<subseteq> S; a \<in> S; b \<in> S; a \<noteq> b; sort_of a = sort_of b\<rbrakk> \<Longrightarrow> P ((a \<rightleftharpoons> b) + p)"
       
  1414   shows "P p"
       
  1415 proof -
       
  1416   have "finite (supp p)" by (simp add: finite_supp)
       
  1417   then show "P p" using S
       
  1418   proof(induct A\<equiv>"supp p" arbitrary: p rule: finite_psubset_induct)
       
  1419     case (psubset p)
       
  1420     then have ih: "\<And>q. supp q \<subset> supp p \<Longrightarrow> P q" by auto
       
  1421     have as: "supp p \<subseteq> S" by fact
       
  1422     { assume "supp p = {}"
       
  1423       then have "p = 0" by (simp add: supp_perm expand_perm_eq)
       
  1424       then have "P p" using zero by simp
       
  1425     }
       
  1426     moreover
       
  1427     { assume "supp p \<noteq> {}"
       
  1428       then obtain a where a0: "a \<in> supp p" by blast
       
  1429       then have a1: "p \<bullet> a \<in> S" "a \<in> S" "sort_of (p \<bullet> a) = sort_of a" "p \<bullet> a \<noteq> a"
       
  1430         using as by (auto simp add: supp_atom supp_perm swap_atom)
       
  1431       let ?q = "(p \<bullet> a \<rightleftharpoons> a) + p"
       
  1432       have a2: "supp ?q \<subseteq> supp p" unfolding supp_perm by (auto simp add: swap_atom)
       
  1433       moreover
       
  1434       have "a \<notin> supp ?q" by (simp add: supp_perm)
       
  1435       then have "supp ?q \<noteq> supp p" using a0 by auto
       
  1436       ultimately have "supp ?q \<subset> supp p" using a2 by auto
       
  1437       then have "P ?q" using ih by simp
       
  1438       moreover
       
  1439       have "supp ?q \<subseteq> S" using as a2 by simp
       
  1440       ultimately  have "P ((p \<bullet> a \<rightleftharpoons> a) + ?q)" using as a1 swap by simp 
       
  1441       moreover 
       
  1442       have "p = (p \<bullet> a \<rightleftharpoons> a) + ?q" by (simp add: expand_perm_eq)
       
  1443       ultimately have "P p" by simp
       
  1444     }
       
  1445     ultimately show "P p" by blast
       
  1446   qed
       
  1447 qed
       
  1448 
       
  1449 lemma perm_simple_struct_induct[case_names zero swap]:
       
  1450   assumes zero: "P 0"
       
  1451   and     swap: "\<And>p a b. \<lbrakk>P p; a \<noteq> b; sort_of a = sort_of b\<rbrakk> \<Longrightarrow> P ((a \<rightleftharpoons> b) + p)"
       
  1452   shows "P p"
       
  1453 by (rule_tac S="supp p" in perm_struct_induct)
       
  1454    (auto intro: zero swap)
       
  1455 
       
  1456 lemma perm_subset_induct[consumes 1, case_names zero swap plus]:
       
  1457   assumes S: "supp p \<subseteq> S"
       
  1458   assumes zero: "P 0"
       
  1459   assumes swap: "\<And>a b. \<lbrakk>sort_of a = sort_of b; a \<noteq> b; a \<in> S; b \<in> S\<rbrakk> \<Longrightarrow> P (a \<rightleftharpoons> b)"
       
  1460   assumes plus: "\<And>p1 p2. \<lbrakk>P p1; P p2; supp p1 \<subseteq> S; supp p2 \<subseteq> S\<rbrakk> \<Longrightarrow> P (p1 + p2)"
       
  1461   shows "P p"
       
  1462 using S
       
  1463 by (induct p rule: perm_struct_induct)
       
  1464    (auto intro: zero plus swap simp add: supp_swap)
       
  1465 
       
  1466 lemma supp_perm_eq:
       
  1467   assumes "(supp x) \<sharp>* p"
       
  1468   shows "p \<bullet> x = x"
       
  1469 proof -
       
  1470   from assms have "supp p \<subseteq> {a. a \<sharp> x}"
       
  1471     unfolding supp_perm fresh_star_def fresh_def by auto
       
  1472   then show "p \<bullet> x = x"
       
  1473   proof (induct p rule: perm_struct_induct)
       
  1474     case zero
       
  1475     show "0 \<bullet> x = x" by simp
       
  1476   next
       
  1477     case (swap p a b)
       
  1478     then have "a \<sharp> x" "b \<sharp> x" "p \<bullet> x = x" by simp_all
       
  1479     then show "((a \<rightleftharpoons> b) + p) \<bullet> x = x" by (simp add: swap_fresh_fresh)
       
  1480   qed
       
  1481 qed
       
  1482 
       
  1483 lemma supp_perm_eq_test:
       
  1484   assumes "(supp x) \<sharp>* p"
       
  1485   shows "p \<bullet> x = x"
       
  1486 proof -
       
  1487   from assms have "supp p \<subseteq> {a. a \<sharp> x}"
       
  1488     unfolding supp_perm fresh_star_def fresh_def by auto
       
  1489   then show "p \<bullet> x = x"
       
  1490   proof (induct p rule: perm_subset_induct)
       
  1491     case zero
       
  1492     show "0 \<bullet> x = x" by simp
       
  1493   next
       
  1494     case (swap a b)
       
  1495     then have "a \<sharp> x" "b \<sharp> x" by simp_all
       
  1496     then show "(a \<rightleftharpoons> b) \<bullet> x = x" by (simp add: swap_fresh_fresh)
       
  1497   next
       
  1498     case (plus p1 p2)
       
  1499     have "p1 \<bullet> x = x" "p2 \<bullet> x = x" by fact+
       
  1500     then show "(p1 + p2) \<bullet> x = x" by simp
       
  1501   qed
       
  1502 qed
       
  1503 
       
  1504 
       
  1505 section {* Avoiding of atom sets *}
       
  1506 
       
  1507 text {* 
       
  1508   For every set of atoms, there is another set of atoms
       
  1509   avoiding a finitely supported c and there is a permutation
       
  1510   which 'translates' between both sets.
       
  1511 *}
       
  1512 
       
  1513 lemma at_set_avoiding_aux:
       
  1514   fixes Xs::"atom set"
       
  1515   and   As::"atom set"
       
  1516   assumes b: "Xs \<subseteq> As"
       
  1517   and     c: "finite As"
       
  1518   shows "\<exists>p. (p \<bullet> Xs) \<inter> As = {} \<and> (supp p) \<subseteq> (Xs \<union> (p \<bullet> Xs))"
       
  1519 proof -
       
  1520   from b c have "finite Xs" by (rule finite_subset)
       
  1521   then show ?thesis using b
       
  1522   proof (induct rule: finite_subset_induct)
       
  1523     case empty
       
  1524     have "0 \<bullet> {} \<inter> As = {}" by simp
       
  1525     moreover
       
  1526     have "supp (0::perm) \<subseteq> {} \<union> 0 \<bullet> {}" by (simp add: supp_zero_perm)
       
  1527     ultimately show ?case by blast
       
  1528   next
       
  1529     case (insert x Xs)
       
  1530     then obtain p where
       
  1531       p1: "(p \<bullet> Xs) \<inter> As = {}" and 
       
  1532       p2: "supp p \<subseteq> (Xs \<union> (p \<bullet> Xs))" by blast
       
  1533     from `x \<in> As` p1 have "x \<notin> p \<bullet> Xs" by fast
       
  1534     with `x \<notin> Xs` p2 have "x \<notin> supp p" by fast
       
  1535     hence px: "p \<bullet> x = x" unfolding supp_perm by simp
       
  1536     have "finite (As \<union> p \<bullet> Xs)"
       
  1537       using `finite As` `finite Xs`
       
  1538       by (simp add: permute_set_eq_image)
       
  1539     then obtain y where "y \<notin> (As \<union> p \<bullet> Xs)" "sort_of y = sort_of x"
       
  1540       by (rule obtain_atom)
       
  1541     hence y: "y \<notin> As" "y \<notin> p \<bullet> Xs" "sort_of y = sort_of x"
       
  1542       by simp_all
       
  1543     let ?q = "(x \<rightleftharpoons> y) + p"
       
  1544     have q: "?q \<bullet> insert x Xs = insert y (p \<bullet> Xs)"
       
  1545       unfolding insert_eqvt
       
  1546       using `p \<bullet> x = x` `sort_of y = sort_of x`
       
  1547       using `x \<notin> p \<bullet> Xs` `y \<notin> p \<bullet> Xs`
       
  1548       by (simp add: swap_atom swap_set_not_in)
       
  1549     have "?q \<bullet> insert x Xs \<inter> As = {}"
       
  1550       using `y \<notin> As` `p \<bullet> Xs \<inter> As = {}`
       
  1551       unfolding q by simp
       
  1552     moreover
       
  1553     have "supp ?q \<subseteq> insert x Xs \<union> ?q \<bullet> insert x Xs"
       
  1554       using p2 unfolding q
       
  1555       by (intro subset_trans [OF supp_plus_perm])
       
  1556          (auto simp add: supp_swap)
       
  1557     ultimately show ?case by blast
       
  1558   qed
       
  1559 qed
       
  1560 
       
  1561 lemma at_set_avoiding:
       
  1562   assumes a: "finite Xs"
       
  1563   and     b: "finite (supp c)"
       
  1564   obtains p::"perm" where "(p \<bullet> Xs)\<sharp>*c" and "(supp p) \<subseteq> (Xs \<union> (p \<bullet> Xs))"
       
  1565   using a b at_set_avoiding_aux [where Xs="Xs" and As="Xs \<union> supp c"]
       
  1566   unfolding fresh_star_def fresh_def by blast
       
  1567 
       
  1568 lemma at_set_avoiding2:
       
  1569   assumes "finite xs"
       
  1570   and     "finite (supp c)" "finite (supp x)"
       
  1571   and     "xs \<sharp>* x"
       
  1572   shows "\<exists>p. (p \<bullet> xs) \<sharp>* c \<and> supp x \<sharp>* p"
       
  1573 using assms
       
  1574 apply(erule_tac c="(c, x)" in at_set_avoiding)
       
  1575 apply(simp add: supp_Pair)
       
  1576 apply(rule_tac x="p" in exI)
       
  1577 apply(simp add: fresh_star_prod)
       
  1578 apply(rule fresh_star_supp_conv)
       
  1579 apply(auto simp add: fresh_star_def)
       
  1580 done
       
  1581 
       
  1582 lemma at_set_avoiding2_atom:
       
  1583   assumes "finite (supp c)" "finite (supp x)"
       
  1584   and     b: "a \<sharp> x"
       
  1585   shows "\<exists>p. (p \<bullet> a) \<sharp> c \<and> supp x \<sharp>* p"
       
  1586 proof -
       
  1587   have a: "{a} \<sharp>* x" unfolding fresh_star_def by (simp add: b)
       
  1588   obtain p where p1: "(p \<bullet> {a}) \<sharp>* c" and p2: "supp x \<sharp>* p"
       
  1589     using at_set_avoiding2[of "{a}" "c" "x"] assms a by blast
       
  1590   have c: "(p \<bullet> a) \<sharp> c" using p1
       
  1591     unfolding fresh_star_def Ball_def 
       
  1592     by(erule_tac x="p \<bullet> a" in allE) (simp add: permute_set_eq)
       
  1593   hence "p \<bullet> a \<sharp> c \<and> supp x \<sharp>* p" using p2 by blast
       
  1594   then show "\<exists>p. (p \<bullet> a) \<sharp> c \<and> supp x \<sharp>* p" by blast
       
  1595 qed
       
  1596 
       
  1597 
       
  1598 section {* Concrete Atoms Types *}
       
  1599 
       
  1600 text {*
       
  1601   Class @{text at_base} allows types containing multiple sorts of atoms.
       
  1602   Class @{text at} only allows types with a single sort.
       
  1603 *}
       
  1604 
       
  1605 class at_base = pt +
       
  1606   fixes atom :: "'a \<Rightarrow> atom"
       
  1607   assumes atom_eq_iff [simp]: "atom a = atom b \<longleftrightarrow> a = b"
       
  1608   assumes atom_eqvt: "p \<bullet> (atom a) = atom (p \<bullet> a)"
       
  1609 
       
  1610 class at = at_base +
       
  1611   assumes sort_of_atom_eq [simp]: "sort_of (atom a) = sort_of (atom b)"
       
  1612 
       
  1613 lemma supp_at_base: 
       
  1614   fixes a::"'a::at_base"
       
  1615   shows "supp a = {atom a}"
       
  1616   by (simp add: supp_atom [symmetric] supp_def atom_eqvt)
       
  1617 
       
  1618 lemma fresh_at_base: 
       
  1619   shows "a \<sharp> b \<longleftrightarrow> a \<noteq> atom b"
       
  1620   unfolding fresh_def by (simp add: supp_at_base)
       
  1621 
       
  1622 instance at_base < fs
       
  1623 proof qed (simp add: supp_at_base)
       
  1624 
       
  1625 lemma at_base_infinite [simp]:
       
  1626   shows "infinite (UNIV :: 'a::at_base set)" (is "infinite ?U")
       
  1627 proof
       
  1628   obtain a :: 'a where "True" by auto
       
  1629   assume "finite ?U"
       
  1630   hence "finite (atom ` ?U)"
       
  1631     by (rule finite_imageI)
       
  1632   then obtain b where b: "b \<notin> atom ` ?U" "sort_of b = sort_of (atom a)"
       
  1633     by (rule obtain_atom)
       
  1634   from b(2) have "b = atom ((atom a \<rightleftharpoons> b) \<bullet> a)"
       
  1635     unfolding atom_eqvt [symmetric]
       
  1636     by (simp add: swap_atom)
       
  1637   hence "b \<in> atom ` ?U" by simp
       
  1638   with b(1) show "False" by simp
       
  1639 qed
       
  1640 
       
  1641 lemma swap_at_base_simps [simp]:
       
  1642   fixes x y::"'a::at_base"
       
  1643   shows "sort_of (atom x) = sort_of (atom y) \<Longrightarrow> (atom x \<rightleftharpoons> atom y) \<bullet> x = y"
       
  1644   and   "sort_of (atom x) = sort_of (atom y) \<Longrightarrow> (atom x \<rightleftharpoons> atom y) \<bullet> y = x"
       
  1645   and   "atom x \<noteq> a \<Longrightarrow> atom x \<noteq> b \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> x = x"
       
  1646   unfolding atom_eq_iff [symmetric]
       
  1647   unfolding atom_eqvt [symmetric]
       
  1648   by simp_all
       
  1649 
       
  1650 lemma obtain_at_base:
       
  1651   assumes X: "finite X"
       
  1652   obtains a::"'a::at_base" where "atom a \<notin> X"
       
  1653 proof -
       
  1654   have "inj (atom :: 'a \<Rightarrow> atom)"
       
  1655     by (simp add: inj_on_def)
       
  1656   with X have "finite (atom -` X :: 'a set)"
       
  1657     by (rule finite_vimageI)
       
  1658   with at_base_infinite have "atom -` X \<noteq> (UNIV :: 'a set)"
       
  1659     by auto
       
  1660   then obtain a :: 'a where "atom a \<notin> X"
       
  1661     by auto
       
  1662   thus ?thesis ..
       
  1663 qed
       
  1664 
       
  1665 lemma supp_finite_set_at_base:
       
  1666   assumes a: "finite S"
       
  1667   shows "supp S = atom ` S"
       
  1668 apply(simp add: supp_of_finite_sets[OF a])
       
  1669 apply(simp add: supp_at_base)
       
  1670 apply(auto)
       
  1671 done
       
  1672 
       
  1673 section {* Infrastructure for concrete atom types *}
       
  1674 
       
  1675 section {* A swapping operation for concrete atoms *}
       
  1676   
       
  1677 definition
       
  1678   flip :: "'a::at_base \<Rightarrow> 'a \<Rightarrow> perm" ("'(_ \<leftrightarrow> _')")
       
  1679 where
       
  1680   "(a \<leftrightarrow> b) = (atom a \<rightleftharpoons> atom b)"
       
  1681 
       
  1682 lemma flip_self [simp]: "(a \<leftrightarrow> a) = 0"
       
  1683   unfolding flip_def by (rule swap_self)
       
  1684 
       
  1685 lemma flip_commute: "(a \<leftrightarrow> b) = (b \<leftrightarrow> a)"
       
  1686   unfolding flip_def by (rule swap_commute)
       
  1687 
       
  1688 lemma minus_flip [simp]: "- (a \<leftrightarrow> b) = (a \<leftrightarrow> b)"
       
  1689   unfolding flip_def by (rule minus_swap)
       
  1690 
       
  1691 lemma add_flip_cancel: "(a \<leftrightarrow> b) + (a \<leftrightarrow> b) = 0"
       
  1692   unfolding flip_def by (rule swap_cancel)
       
  1693 
       
  1694 lemma permute_flip_cancel [simp]: "(a \<leftrightarrow> b) \<bullet> (a \<leftrightarrow> b) \<bullet> x = x"
       
  1695   unfolding permute_plus [symmetric] add_flip_cancel by simp
       
  1696 
       
  1697 lemma permute_flip_cancel2 [simp]: "(a \<leftrightarrow> b) \<bullet> (b \<leftrightarrow> a) \<bullet> x = x"
       
  1698   by (simp add: flip_commute)
       
  1699 
       
  1700 lemma flip_eqvt: 
       
  1701   fixes a b c::"'a::at_base"
       
  1702   shows "p \<bullet> (a \<leftrightarrow> b) = (p \<bullet> a \<leftrightarrow> p \<bullet> b)"
       
  1703   unfolding flip_def
       
  1704   by (simp add: swap_eqvt atom_eqvt)
       
  1705 
       
  1706 lemma flip_at_base_simps [simp]:
       
  1707   shows "sort_of (atom a) = sort_of (atom b) \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> a = b"
       
  1708   and   "sort_of (atom a) = sort_of (atom b) \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> b = a"
       
  1709   and   "\<lbrakk>a \<noteq> c; b \<noteq> c\<rbrakk> \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> c = c"
       
  1710   and   "sort_of (atom a) \<noteq> sort_of (atom b) \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> x = x"
       
  1711   unfolding flip_def
       
  1712   unfolding atom_eq_iff [symmetric]
       
  1713   unfolding atom_eqvt [symmetric]
       
  1714   by simp_all
       
  1715 
       
  1716 text {* the following two lemmas do not hold for at_base, 
       
  1717   only for single sort atoms from at *}
       
  1718 
       
  1719 lemma permute_flip_at:
       
  1720   fixes a b c::"'a::at"
       
  1721   shows "(a \<leftrightarrow> b) \<bullet> c = (if c = a then b else if c = b then a else c)"
       
  1722   unfolding flip_def
       
  1723   apply (rule atom_eq_iff [THEN iffD1])
       
  1724   apply (subst atom_eqvt [symmetric])
       
  1725   apply (simp add: swap_atom)
       
  1726   done
       
  1727 
       
  1728 lemma flip_at_simps [simp]:
       
  1729   fixes a b::"'a::at"
       
  1730   shows "(a \<leftrightarrow> b) \<bullet> a = b" 
       
  1731   and   "(a \<leftrightarrow> b) \<bullet> b = a"
       
  1732   unfolding permute_flip_at by simp_all
       
  1733 
       
  1734 lemma flip_fresh_fresh:
       
  1735   fixes a b::"'a::at_base"
       
  1736   assumes "atom a \<sharp> x" "atom b \<sharp> x"
       
  1737   shows "(a \<leftrightarrow> b) \<bullet> x = x"
       
  1738 using assms
       
  1739 by (simp add: flip_def swap_fresh_fresh)
       
  1740 
       
  1741 subsection {* Syntax for coercing at-elements to the atom-type *}
       
  1742 
       
  1743 syntax
       
  1744   "_atom_constrain" :: "logic \<Rightarrow> type \<Rightarrow> logic" ("_:::_" [4, 0] 3)
       
  1745 
       
  1746 translations
       
  1747   "_atom_constrain a t" => "CONST atom (_constrain a t)"
       
  1748 
       
  1749 
       
  1750 subsection {* A lemma for proving instances of class @{text at}. *}
       
  1751 
       
  1752 setup {* Sign.add_const_constraint (@{const_name "permute"}, NONE) *}
       
  1753 setup {* Sign.add_const_constraint (@{const_name "atom"}, NONE) *}
       
  1754 
       
  1755 text {*
       
  1756   New atom types are defined as subtypes of @{typ atom}.
       
  1757 *}
       
  1758 
       
  1759 lemma exists_eq_simple_sort: 
       
  1760   shows "\<exists>a. a \<in> {a. sort_of a = s}"
       
  1761   by (rule_tac x="Atom s 0" in exI, simp)
       
  1762 
       
  1763 lemma exists_eq_sort: 
       
  1764   shows "\<exists>a. a \<in> {a. sort_of a \<in> range sort_fun}"
       
  1765   by (rule_tac x="Atom (sort_fun x) y" in exI, simp)
       
  1766 
       
  1767 lemma at_base_class:
       
  1768   fixes sort_fun :: "'b \<Rightarrow>atom_sort"
       
  1769   fixes Rep :: "'a \<Rightarrow> atom" and Abs :: "atom \<Rightarrow> 'a"
       
  1770   assumes type: "type_definition Rep Abs {a. sort_of a \<in> range sort_fun}"
       
  1771   assumes atom_def: "\<And>a. atom a = Rep a"
       
  1772   assumes permute_def: "\<And>p a. p \<bullet> a = Abs (p \<bullet> Rep a)"
       
  1773   shows "OFCLASS('a, at_base_class)"
       
  1774 proof
       
  1775   interpret type_definition Rep Abs "{a. sort_of a \<in> range sort_fun}" by (rule type)
       
  1776   have sort_of_Rep: "\<And>a. sort_of (Rep a) \<in> range sort_fun" using Rep by simp
       
  1777   fix a b :: 'a and p p1 p2 :: perm
       
  1778   show "0 \<bullet> a = a"
       
  1779     unfolding permute_def by (simp add: Rep_inverse)
       
  1780   show "(p1 + p2) \<bullet> a = p1 \<bullet> p2 \<bullet> a"
       
  1781     unfolding permute_def by (simp add: Abs_inverse sort_of_Rep)
       
  1782   show "atom a = atom b \<longleftrightarrow> a = b"
       
  1783     unfolding atom_def by (simp add: Rep_inject)
       
  1784   show "p \<bullet> atom a = atom (p \<bullet> a)"
       
  1785     unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep)
       
  1786 qed
       
  1787 
       
  1788 (*
       
  1789 lemma at_class:
       
  1790   fixes s :: atom_sort
       
  1791   fixes Rep :: "'a \<Rightarrow> atom" and Abs :: "atom \<Rightarrow> 'a"
       
  1792   assumes type: "type_definition Rep Abs {a. sort_of a \<in> range (\<lambda>x::unit. s)}"
       
  1793   assumes atom_def: "\<And>a. atom a = Rep a"
       
  1794   assumes permute_def: "\<And>p a. p \<bullet> a = Abs (p \<bullet> Rep a)"
       
  1795   shows "OFCLASS('a, at_class)"
       
  1796 proof
       
  1797   interpret type_definition Rep Abs "{a. sort_of a \<in> range (\<lambda>x::unit. s)}" by (rule type)
       
  1798   have sort_of_Rep: "\<And>a. sort_of (Rep a) = s" using Rep by (simp add: image_def)
       
  1799   fix a b :: 'a and p p1 p2 :: perm
       
  1800   show "0 \<bullet> a = a"
       
  1801     unfolding permute_def by (simp add: Rep_inverse)
       
  1802   show "(p1 + p2) \<bullet> a = p1 \<bullet> p2 \<bullet> a"
       
  1803     unfolding permute_def by (simp add: Abs_inverse sort_of_Rep)
       
  1804   show "sort_of (atom a) = sort_of (atom b)"
       
  1805     unfolding atom_def by (simp add: sort_of_Rep)
       
  1806   show "atom a = atom b \<longleftrightarrow> a = b"
       
  1807     unfolding atom_def by (simp add: Rep_inject)
       
  1808   show "p \<bullet> atom a = atom (p \<bullet> a)"
       
  1809     unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep)
       
  1810 qed
       
  1811 *)
       
  1812 
       
  1813 lemma at_class:
       
  1814   fixes s :: atom_sort
       
  1815   fixes Rep :: "'a \<Rightarrow> atom" and Abs :: "atom \<Rightarrow> 'a"
       
  1816   assumes type: "type_definition Rep Abs {a. sort_of a = s}"
       
  1817   assumes atom_def: "\<And>a. atom a = Rep a"
       
  1818   assumes permute_def: "\<And>p a. p \<bullet> a = Abs (p \<bullet> Rep a)"
       
  1819   shows "OFCLASS('a, at_class)"
       
  1820 proof
       
  1821   interpret type_definition Rep Abs "{a. sort_of a = s}" by (rule type)
       
  1822   have sort_of_Rep: "\<And>a. sort_of (Rep a) = s" using Rep by (simp add: image_def)
       
  1823   fix a b :: 'a and p p1 p2 :: perm
       
  1824   show "0 \<bullet> a = a"
       
  1825     unfolding permute_def by (simp add: Rep_inverse)
       
  1826   show "(p1 + p2) \<bullet> a = p1 \<bullet> p2 \<bullet> a"
       
  1827     unfolding permute_def by (simp add: Abs_inverse sort_of_Rep)
       
  1828   show "sort_of (atom a) = sort_of (atom b)"
       
  1829     unfolding atom_def by (simp add: sort_of_Rep)
       
  1830   show "atom a = atom b \<longleftrightarrow> a = b"
       
  1831     unfolding atom_def by (simp add: Rep_inject)
       
  1832   show "p \<bullet> atom a = atom (p \<bullet> a)"
       
  1833     unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep)
       
  1834 qed
       
  1835 
       
  1836 setup {* Sign.add_const_constraint
       
  1837   (@{const_name "permute"}, SOME @{typ "perm \<Rightarrow> 'a::pt \<Rightarrow> 'a"}) *}
       
  1838 setup {* Sign.add_const_constraint
       
  1839   (@{const_name "atom"}, SOME @{typ "'a::at_base \<Rightarrow> atom"}) *}
       
  1840 
       
  1841 
       
  1842 
       
  1843 section {* The freshness lemma according to Andy Pitts *}
       
  1844 
       
  1845 lemma freshness_lemma:
       
  1846   fixes h :: "'a::at \<Rightarrow> 'b::pt"
       
  1847   assumes a: "\<exists>a. atom a \<sharp> (h, h a)"
       
  1848   shows  "\<exists>x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x"
       
  1849 proof -
       
  1850   from a obtain b where a1: "atom b \<sharp> h" and a2: "atom b \<sharp> h b"
       
  1851     by (auto simp add: fresh_Pair)
       
  1852   show "\<exists>x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x"
       
  1853   proof (intro exI allI impI)
       
  1854     fix a :: 'a
       
  1855     assume a3: "atom a \<sharp> h"
       
  1856     show "h a = h b"
       
  1857     proof (cases "a = b")
       
  1858       assume "a = b"
       
  1859       thus "h a = h b" by simp
       
  1860     next
       
  1861       assume "a \<noteq> b"
       
  1862       hence "atom a \<sharp> b" by (simp add: fresh_at_base)
       
  1863       with a3 have "atom a \<sharp> h b" 
       
  1864         by (rule fresh_fun_app)
       
  1865       with a2 have d1: "(atom b \<rightleftharpoons> atom a) \<bullet> (h b) = (h b)"
       
  1866         by (rule swap_fresh_fresh)
       
  1867       from a1 a3 have d2: "(atom b \<rightleftharpoons> atom a) \<bullet> h = h"
       
  1868         by (rule swap_fresh_fresh)
       
  1869       from d1 have "h b = (atom b \<rightleftharpoons> atom a) \<bullet> (h b)" by simp
       
  1870       also have "\<dots> = ((atom b \<rightleftharpoons> atom a) \<bullet> h) ((atom b \<rightleftharpoons> atom a) \<bullet> b)"
       
  1871         by (rule permute_fun_app_eq)
       
  1872       also have "\<dots> = h a"
       
  1873         using d2 by simp
       
  1874       finally show "h a = h b"  by simp
       
  1875     qed
       
  1876   qed
       
  1877 qed
       
  1878 
       
  1879 lemma freshness_lemma_unique:
       
  1880   fixes h :: "'a::at \<Rightarrow> 'b::pt"
       
  1881   assumes a: "\<exists>a. atom a \<sharp> (h, h a)"
       
  1882   shows "\<exists>!x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x"
       
  1883 proof (rule ex_ex1I)
       
  1884   from a show "\<exists>x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x"
       
  1885     by (rule freshness_lemma)
       
  1886 next
       
  1887   fix x y
       
  1888   assume x: "\<forall>a. atom a \<sharp> h \<longrightarrow> h a = x"
       
  1889   assume y: "\<forall>a. atom a \<sharp> h \<longrightarrow> h a = y"
       
  1890   from a x y show "x = y"
       
  1891     by (auto simp add: fresh_Pair)
       
  1892 qed
       
  1893 
       
  1894 text {* packaging the freshness lemma into a function *}
       
  1895 
       
  1896 definition
       
  1897   fresh_fun :: "('a::at \<Rightarrow> 'b::pt) \<Rightarrow> 'b"
       
  1898 where
       
  1899   "fresh_fun h = (THE x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x)"
       
  1900 
       
  1901 lemma fresh_fun_apply:
       
  1902   fixes h :: "'a::at \<Rightarrow> 'b::pt"
       
  1903   assumes a: "\<exists>a. atom a \<sharp> (h, h a)"
       
  1904   assumes b: "atom a \<sharp> h"
       
  1905   shows "fresh_fun h = h a"
       
  1906 unfolding fresh_fun_def
       
  1907 proof (rule the_equality)
       
  1908   show "\<forall>a'. atom a' \<sharp> h \<longrightarrow> h a' = h a"
       
  1909   proof (intro strip)
       
  1910     fix a':: 'a
       
  1911     assume c: "atom a' \<sharp> h"
       
  1912     from a have "\<exists>x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x" by (rule freshness_lemma)
       
  1913     with b c show "h a' = h a" by auto
       
  1914   qed
       
  1915 next
       
  1916   fix fr :: 'b
       
  1917   assume "\<forall>a. atom a \<sharp> h \<longrightarrow> h a = fr"
       
  1918   with b show "fr = h a" by auto
       
  1919 qed
       
  1920 
       
  1921 lemma fresh_fun_apply':
       
  1922   fixes h :: "'a::at \<Rightarrow> 'b::pt"
       
  1923   assumes a: "atom a \<sharp> h" "atom a \<sharp> h a"
       
  1924   shows "fresh_fun h = h a"
       
  1925   apply (rule fresh_fun_apply)
       
  1926   apply (auto simp add: fresh_Pair intro: a)
       
  1927   done
       
  1928 
       
  1929 lemma fresh_fun_eqvt:
       
  1930   fixes h :: "'a::at \<Rightarrow> 'b::pt"
       
  1931   assumes a: "\<exists>a. atom a \<sharp> (h, h a)"
       
  1932   shows "p \<bullet> (fresh_fun h) = fresh_fun (p \<bullet> h)"
       
  1933   using a
       
  1934   apply (clarsimp simp add: fresh_Pair)
       
  1935   apply (subst fresh_fun_apply', assumption+)
       
  1936   apply (drule fresh_permute_iff [where p=p, THEN iffD2])
       
  1937   apply (drule fresh_permute_iff [where p=p, THEN iffD2])
       
  1938   apply (simp add: atom_eqvt permute_fun_app_eq [where f=h])
       
  1939   apply (erule (1) fresh_fun_apply' [symmetric])
       
  1940   done
       
  1941 
       
  1942 lemma fresh_fun_supports:
       
  1943   fixes h :: "'a::at \<Rightarrow> 'b::pt"
       
  1944   assumes a: "\<exists>a. atom a \<sharp> (h, h a)"
       
  1945   shows "(supp h) supports (fresh_fun h)"
       
  1946   apply (simp add: supports_def fresh_def [symmetric])
       
  1947   apply (simp add: fresh_fun_eqvt [OF a] swap_fresh_fresh)
       
  1948   done
       
  1949 
       
  1950 notation fresh_fun (binder "FRESH " 10)
       
  1951 
       
  1952 lemma FRESH_f_iff:
       
  1953   fixes P :: "'a::at \<Rightarrow> 'b::pure"
       
  1954   fixes f :: "'b \<Rightarrow> 'c::pure"
       
  1955   assumes P: "finite (supp P)"
       
  1956   shows "(FRESH x. f (P x)) = f (FRESH x. P x)"
       
  1957 proof -
       
  1958   obtain a::'a where "atom a \<notin> supp P"
       
  1959     using P by (rule obtain_at_base)
       
  1960   hence "atom a \<sharp> P"
       
  1961     by (simp add: fresh_def)
       
  1962   show "(FRESH x. f (P x)) = f (FRESH x. P x)"
       
  1963     apply (subst fresh_fun_apply' [where a=a, OF _ pure_fresh])
       
  1964     apply (cut_tac `atom a \<sharp> P`)
       
  1965     apply (simp add: fresh_conv_MOST)
       
  1966     apply (elim MOST_rev_mp, rule MOST_I, clarify)
       
  1967     apply (simp add: permute_fun_def permute_pure fun_eq_iff)
       
  1968     apply (subst fresh_fun_apply' [where a=a, OF `atom a \<sharp> P` pure_fresh])
       
  1969     apply (rule refl)
       
  1970     done
       
  1971 qed
       
  1972 
       
  1973 lemma FRESH_binop_iff:
       
  1974   fixes P :: "'a::at \<Rightarrow> 'b::pure"
       
  1975   fixes Q :: "'a::at \<Rightarrow> 'c::pure"
       
  1976   fixes binop :: "'b \<Rightarrow> 'c \<Rightarrow> 'd::pure"
       
  1977   assumes P: "finite (supp P)" 
       
  1978   and     Q: "finite (supp Q)"
       
  1979   shows "(FRESH x. binop (P x) (Q x)) = binop (FRESH x. P x) (FRESH x. Q x)"
       
  1980 proof -
       
  1981   from assms have "finite (supp P \<union> supp Q)" by simp
       
  1982   then obtain a::'a where "atom a \<notin> (supp P \<union> supp Q)"
       
  1983     by (rule obtain_at_base)
       
  1984   hence "atom a \<sharp> P" and "atom a \<sharp> Q"
       
  1985     by (simp_all add: fresh_def)
       
  1986   show ?thesis
       
  1987     apply (subst fresh_fun_apply' [where a=a, OF _ pure_fresh])
       
  1988     apply (cut_tac `atom a \<sharp> P` `atom a \<sharp> Q`)
       
  1989     apply (simp add: fresh_conv_MOST)
       
  1990     apply (elim MOST_rev_mp, rule MOST_I, clarify)
       
  1991     apply (simp add: permute_fun_def permute_pure fun_eq_iff)
       
  1992     apply (subst fresh_fun_apply' [where a=a, OF `atom a \<sharp> P` pure_fresh])
       
  1993     apply (subst fresh_fun_apply' [where a=a, OF `atom a \<sharp> Q` pure_fresh])
       
  1994     apply (rule refl)
       
  1995     done
       
  1996 qed
       
  1997 
       
  1998 lemma FRESH_conj_iff:
       
  1999   fixes P Q :: "'a::at \<Rightarrow> bool"
       
  2000   assumes P: "finite (supp P)" and Q: "finite (supp Q)"
       
  2001   shows "(FRESH x. P x \<and> Q x) \<longleftrightarrow> (FRESH x. P x) \<and> (FRESH x. Q x)"
       
  2002 using P Q by (rule FRESH_binop_iff)
       
  2003 
       
  2004 lemma FRESH_disj_iff:
       
  2005   fixes P Q :: "'a::at \<Rightarrow> bool"
       
  2006   assumes P: "finite (supp P)" and Q: "finite (supp Q)"
       
  2007   shows "(FRESH x. P x \<or> Q x) \<longleftrightarrow> (FRESH x. P x) \<or> (FRESH x. Q x)"
       
  2008 using P Q by (rule FRESH_binop_iff)
       
  2009 
       
  2010 
       
  2011 section {* Library functions for the nominal infrastructure *}
       
  2012 
       
  2013 use "nominal_library.ML"
       
  2014 
       
  2015 
       
  2016 section {* Automation for creating concrete atom types *}
       
  2017 
       
  2018 text {* at the moment only single-sort concrete atoms are supported *}
       
  2019 
       
  2020 use "nominal_atoms.ML"
       
  2021 
       
  2022 
       
  2023 
       
  2024 end