| 24895 |      1 | (* $Id$ *)
 | 
|  |      2 | 
 | 
|  |      3 | theory Support 
 | 
| 25751 |      4 |   imports "../Nominal" 
 | 
| 24895 |      5 | begin
 | 
|  |      6 | 
 | 
|  |      7 | text {* 
 | 
|  |      8 |   An example showing that in general
 | 
|  |      9 | 
 | 
|  |     10 |   x\<sharp>(A \<union> B) does not imply  x\<sharp>A and  x\<sharp>B
 | 
|  |     11 | 
 | 
| 26262 |     12 |   For this we set A to the set of even atoms and B to 
 | 
|  |     13 |   the set of odd atoms. Then A \<union> B, that is the set of 
 | 
|  |     14 |   all atoms, has empty support. The sets A, respectively B, 
 | 
|  |     15 |   however have the set of all atoms as their support. 
 | 
| 24895 |     16 | *}
 | 
|  |     17 | 
 | 
|  |     18 | atom_decl atom
 | 
|  |     19 | 
 | 
| 25751 |     20 | text {* The set of even atoms. *}
 | 
| 24895 |     21 | abbreviation
 | 
|  |     22 |   EVEN :: "atom set"
 | 
|  |     23 | where
 | 
|  |     24 |   "EVEN \<equiv> {atom n | n. \<exists>i. n=2*i}"
 | 
|  |     25 | 
 | 
| 25751 |     26 | text {* The set of odd atoms: *}
 | 
| 24895 |     27 | abbreviation  
 | 
|  |     28 |   ODD :: "atom set"
 | 
|  |     29 | where
 | 
|  |     30 |   "ODD \<equiv> {atom n | n. \<exists>i. n=2*i+1}"
 | 
|  |     31 | 
 | 
| 25751 |     32 | text {* An atom is either even or odd. *}
 | 
| 24895 |     33 | lemma even_or_odd:
 | 
| 37388 |     34 |   fixes n :: nat
 | 
| 24895 |     35 |   shows "\<exists>i. (n = 2*i) \<or> (n=2*i+1)"
 | 
|  |     36 |   by (induct n) (presburger)+
 | 
|  |     37 | 
 | 
| 26055 |     38 | text {* 
 | 
|  |     39 |   The union of even and odd atoms is the set of all atoms. 
 | 
| 26262 |     40 |   (Unfortunately I do not know a simpler proof of this fact.) *}
 | 
| 24895 |     41 | lemma EVEN_union_ODD:
 | 
|  |     42 |   shows "EVEN \<union> ODD = UNIV"
 | 
| 25139 |     43 |   using even_or_odd
 | 
| 24895 |     44 | proof -
 | 
|  |     45 |   have "EVEN \<union> ODD = (\<lambda>n. atom n) ` {n. \<exists>i. n = 2*i} \<union> (\<lambda>n. atom n) ` {n. \<exists>i. n = 2*i+1}" by auto
 | 
|  |     46 |   also have "\<dots> = (\<lambda>n. atom n) ` ({n. \<exists>i. n = 2*i} \<union> {n. \<exists>i. n = 2*i+1})" by auto
 | 
|  |     47 |   also have "\<dots> = (\<lambda>n. atom n) ` ({n. \<exists>i. n = 2*i \<or> n = 2*i+1})" by auto
 | 
|  |     48 |   also have "\<dots> = (\<lambda>n. atom n) ` (UNIV::nat set)" using even_or_odd by auto
 | 
|  |     49 |   also have "\<dots> = (UNIV::atom set)" using atom.exhaust
 | 
| 40702 |     50 |     by (auto simp add: surj_def)
 | 
| 24895 |     51 |   finally show "EVEN \<union> ODD = UNIV" by simp
 | 
|  |     52 | qed
 | 
|  |     53 | 
 | 
| 25751 |     54 | text {* The sets of even and odd atoms are disjunct. *}
 | 
| 24895 |     55 | lemma EVEN_intersect_ODD:
 | 
|  |     56 |   shows "EVEN \<inter> ODD = {}"
 | 
|  |     57 |   using even_or_odd
 | 
|  |     58 |   by (auto) (presburger)
 | 
|  |     59 | 
 | 
| 25751 |     60 | text {* 
 | 
|  |     61 |   The preceeding two lemmas help us to prove 
 | 
| 26055 |     62 |   the following two useful equalities: *}
 | 
|  |     63 | 
 | 
| 24895 |     64 | lemma UNIV_subtract:
 | 
|  |     65 |   shows "UNIV - EVEN = ODD"
 | 
|  |     66 |   and   "UNIV - ODD  = EVEN"
 | 
|  |     67 |   using EVEN_union_ODD EVEN_intersect_ODD
 | 
|  |     68 |   by (blast)+
 | 
|  |     69 | 
 | 
| 25751 |     70 | text {* The sets EVEN and ODD are infinite. *}
 | 
| 24895 |     71 | lemma EVEN_ODD_infinite:
 | 
|  |     72 |   shows "infinite EVEN"
 | 
|  |     73 |   and   "infinite ODD"
 | 
| 25751 |     74 | unfolding infinite_iff_countable_subset
 | 
|  |     75 | proof -
 | 
|  |     76 |   let ?f = "\<lambda>n. atom (2*n)"
 | 
|  |     77 |   have "inj ?f \<and> range ?f \<subseteq> EVEN" by (auto simp add: inj_on_def)
 | 
|  |     78 |   then show "\<exists>f::nat\<Rightarrow>atom. inj f \<and> range f \<subseteq> EVEN" by (rule_tac exI)
 | 
|  |     79 | next
 | 
|  |     80 |   let ?f = "\<lambda>n. atom (2*n+1)"
 | 
|  |     81 |   have "inj ?f \<and> range ?f \<subseteq> ODD" by (auto simp add: inj_on_def)
 | 
|  |     82 |   then show "\<exists>f::nat\<Rightarrow>atom. inj f \<and> range f \<subseteq> ODD" by (rule_tac exI)
 | 
|  |     83 | qed
 | 
| 24895 |     84 | 
 | 
| 25139 |     85 | text {* 
 | 
| 26262 |     86 |   A general fact about a set S of atoms that is both infinite and 
 | 
| 27360 |     87 |   coinfinite. Then S has all atoms as its support. Steve Zdancewic 
 | 
| 26055 |     88 |   helped with proving this fact. *}
 | 
|  |     89 | 
 | 
| 24895 |     90 | lemma supp_infinite_coinfinite:
 | 
|  |     91 |   fixes S::"atom set"
 | 
| 25867 |     92 |   assumes asm1: "infinite S"
 | 
|  |     93 |   and     asm2: "infinite (UNIV-S)"
 | 
| 24895 |     94 |   shows "(supp S) = (UNIV::atom set)"
 | 
|  |     95 | proof -
 | 
|  |     96 |   have "\<forall>(x::atom). x\<in>(supp S)"
 | 
|  |     97 |   proof
 | 
|  |     98 |     fix x::"atom"
 | 
|  |     99 |     show "x\<in>(supp S)"
 | 
|  |    100 |     proof (cases "x\<in>S")
 | 
|  |    101 |       case True
 | 
|  |    102 |       have "x\<in>S" by fact
 | 
| 26806 |    103 |       hence "\<forall>b\<in>(UNIV-S). [(x,b)]\<bullet>S\<noteq>S" by (auto simp add: perm_set_eq calc_atm)
 | 
| 25867 |    104 |       with asm2 have "infinite {b\<in>(UNIV-S). [(x,b)]\<bullet>S\<noteq>S}" by (rule infinite_Collection)
 | 
| 24895 |    105 |       hence "infinite {b. [(x,b)]\<bullet>S\<noteq>S}" by (rule_tac infinite_super, auto)
 | 
|  |    106 |       then show "x\<in>(supp S)" by (simp add: supp_def)
 | 
|  |    107 |     next
 | 
|  |    108 |       case False
 | 
|  |    109 |       have "x\<notin>S" by fact
 | 
| 26806 |    110 |       hence "\<forall>b\<in>S. [(x,b)]\<bullet>S\<noteq>S" by (auto simp add: perm_set_eq calc_atm)
 | 
| 25867 |    111 |       with asm1 have "infinite {b\<in>S. [(x,b)]\<bullet>S\<noteq>S}" by (rule infinite_Collection)
 | 
| 24895 |    112 |       hence "infinite {b. [(x,b)]\<bullet>S\<noteq>S}" by (rule_tac infinite_super, auto)
 | 
|  |    113 |       then show "x\<in>(supp S)" by (simp add: supp_def)
 | 
|  |    114 |     qed
 | 
|  |    115 |   qed
 | 
|  |    116 |   then show "(supp S) = (UNIV::atom set)" by auto
 | 
|  |    117 | qed
 | 
|  |    118 | 
 | 
| 25751 |    119 | text {* As a corollary we get that EVEN and ODD have infinite support. *}
 | 
| 24895 |    120 | lemma EVEN_ODD_supp:
 | 
|  |    121 |   shows "supp EVEN = (UNIV::atom set)"
 | 
|  |    122 |   and   "supp ODD  = (UNIV::atom set)"
 | 
|  |    123 |   using supp_infinite_coinfinite UNIV_subtract EVEN_ODD_infinite
 | 
|  |    124 |   by simp_all
 | 
|  |    125 | 
 | 
| 25751 |    126 | text {* 
 | 
|  |    127 |   The set of all atoms has empty support, since any swappings leaves 
 | 
| 26055 |    128 |   this set unchanged. *}
 | 
|  |    129 | 
 | 
| 24895 |    130 | lemma UNIV_supp:
 | 
|  |    131 |   shows "supp (UNIV::atom set) = ({}::atom set)"
 | 
|  |    132 | proof -
 | 
|  |    133 |   have "\<forall>(x::atom) (y::atom). [(x,y)]\<bullet>UNIV = (UNIV::atom set)"
 | 
| 26806 |    134 |     by (auto simp add: perm_set_eq calc_atm)
 | 
| 25751 |    135 |   then show "supp (UNIV::atom set) = ({}::atom set)" by (simp add: supp_def)
 | 
| 24895 |    136 | qed
 | 
|  |    137 | 
 | 
| 25751 |    138 | text {* Putting everything together. *}
 | 
| 24895 |    139 | theorem EVEN_ODD_freshness:
 | 
|  |    140 |   fixes x::"atom"
 | 
|  |    141 |   shows "x\<sharp>(EVEN \<union> ODD)"
 | 
|  |    142 |   and   "\<not>x\<sharp>EVEN"
 | 
|  |    143 |   and   "\<not>x\<sharp>ODD"
 | 
|  |    144 |   by (auto simp only: fresh_def EVEN_union_ODD EVEN_ODD_supp UNIV_supp)
 | 
|  |    145 | 
 | 
| 26055 |    146 | text {* Moral: support is a sublte notion. *}
 | 
|  |    147 | 
 | 
| 24895 |    148 | end |