Nominal-General/Nominal2_Atoms.thy
author Christian Urban <urbanc@in.tum.de>
Sat, 24 Apr 2010 09:49:23 +0200
changeset 1941 d33781f9d2c7
parent 1933 9eab1dfc14d2
child 1962 84a13d1e2511
permissions -rw-r--r--
merged
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
     1
(*  Title:      Nominal2_Atoms
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
     2
    Authors:    Brian Huffman, Christian Urban
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
     3
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
     4
    Definitions for concrete atom types. 
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
     5
*)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
     6
theory Nominal2_Atoms
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
     7
imports Nominal2_Base
1933
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
     8
        Nominal2_Eqvt
1079
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
     9
uses ("nominal_atoms.ML")
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    10
begin
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    11
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    12
section {* Concrete atom types *}
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    13
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    14
text {*
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    15
  Class @{text at_base} allows types containing multiple sorts of atoms.
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    16
  Class @{text at} only allows types with a single sort.
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    17
*}
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    18
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    19
class at_base = pt +
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    20
  fixes atom :: "'a \<Rightarrow> atom"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    21
  assumes atom_eq_iff [simp]: "atom a = atom b \<longleftrightarrow> a = b"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    22
  assumes atom_eqvt: "p \<bullet> (atom a) = atom (p \<bullet> a)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    23
1933
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
    24
declare atom_eqvt[eqvt]
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
    25
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    26
class at = at_base +
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    27
  assumes sort_of_atom_eq [simp]: "sort_of (atom a) = sort_of (atom b)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    28
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    29
lemma supp_at_base: 
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    30
  fixes a::"'a::at_base"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    31
  shows "supp a = {atom a}"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    32
  by (simp add: supp_atom [symmetric] supp_def atom_eqvt)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    33
1079
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
    34
lemma fresh_at_base: 
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    35
  shows "a \<sharp> b \<longleftrightarrow> a \<noteq> atom b"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    36
  unfolding fresh_def by (simp add: supp_at_base)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    37
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    38
instance at_base < fs
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    39
proof qed (simp add: supp_at_base)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    40
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    41
lemma at_base_infinite [simp]:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    42
  shows "infinite (UNIV :: 'a::at_base set)" (is "infinite ?U")
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    43
proof
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    44
  obtain a :: 'a where "True" by auto
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    45
  assume "finite ?U"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    46
  hence "finite (atom ` ?U)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    47
    by (rule finite_imageI)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    48
  then obtain b where b: "b \<notin> atom ` ?U" "sort_of b = sort_of (atom a)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    49
    by (rule obtain_atom)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    50
  from b(2) have "b = atom ((atom a \<rightleftharpoons> b) \<bullet> a)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    51
    unfolding atom_eqvt [symmetric]
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    52
    by (simp add: swap_atom)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    53
  hence "b \<in> atom ` ?U" by simp
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    54
  with b(1) show "False" by simp
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    55
qed
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    56
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    57
lemma swap_at_base_simps [simp]:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    58
  fixes x y::"'a::at_base"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    59
  shows "sort_of (atom x) = sort_of (atom y) \<Longrightarrow> (atom x \<rightleftharpoons> atom y) \<bullet> x = y"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    60
  and   "sort_of (atom x) = sort_of (atom y) \<Longrightarrow> (atom x \<rightleftharpoons> atom y) \<bullet> y = x"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    61
  and   "atom x \<noteq> a \<Longrightarrow> atom x \<noteq> b \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> x = x"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    62
  unfolding atom_eq_iff [symmetric]
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    63
  unfolding atom_eqvt [symmetric]
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    64
  by simp_all
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    65
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    66
lemma obtain_at_base:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    67
  assumes X: "finite X"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    68
  obtains a::"'a::at_base" where "atom a \<notin> X"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    69
proof -
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    70
  have "inj (atom :: 'a \<Rightarrow> atom)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    71
    by (simp add: inj_on_def)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    72
  with X have "finite (atom -` X :: 'a set)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    73
    by (rule finite_vimageI)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    74
  with at_base_infinite have "atom -` X \<noteq> (UNIV :: 'a set)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    75
    by auto
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    76
  then obtain a :: 'a where "atom a \<notin> X"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    77
    by auto
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    78
  thus ?thesis ..
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    79
qed
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
    80
1933
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
    81
lemma atom_image_cong:
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
    82
  fixes X Y::"('a::at_base) set"
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
    83
  shows "(atom ` X = atom ` Y) = (X = Y)"
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
    84
  apply(rule inj_image_eq_iff)
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
    85
  apply(simp add: inj_on_def)
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
    86
  done
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
    87
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
    88
lemma atom_image_supp:
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
    89
  "supp S = supp (atom ` S)"
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
    90
  apply(simp add: supp_def)
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
    91
  apply(simp add: image_eqvt)
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
    92
  apply(subst (2) permute_fun_def)
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
    93
  apply(simp add: atom_eqvt)
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
    94
  apply(simp add: atom_image_cong)
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
    95
  done
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
    96
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
    97
lemma supp_finite_at_set:
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
    98
  fixes S::"('a::at) set"
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
    99
  assumes a: "finite S"
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   100
  shows "supp S = atom ` S"
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   101
  apply(rule finite_supp_unique)
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   102
  apply(simp add: supports_def)
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   103
  apply(rule allI)+
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   104
  apply(rule impI)
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   105
  apply(rule swap_fresh_fresh)
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   106
  apply(simp add: fresh_def)
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   107
  apply(simp add: atom_image_supp)
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   108
  apply(subst supp_finite_atom_set)
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   109
  apply(rule finite_imageI)
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   110
  apply(simp add: a)
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   111
  apply(simp)
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   112
  apply(simp add: fresh_def)
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   113
  apply(simp add: atom_image_supp)
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   114
  apply(subst supp_finite_atom_set)
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   115
  apply(rule finite_imageI)
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   116
  apply(simp add: a)
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   117
  apply(simp)
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   118
  apply(rule finite_imageI)
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   119
  apply(simp add: a)
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   120
  apply(drule swap_set_in)
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   121
  apply(assumption)
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   122
  apply(simp)
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   123
  apply(simp add: image_eqvt)
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   124
  apply(simp add: permute_fun_def)
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   125
  apply(simp add: atom_eqvt)
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   126
  apply(simp add: atom_image_cong)
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   127
  done
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   128
1941
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   129
lemma supp_finite_at_set_aux:
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   130
  fixes S::"('a::at) set"
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   131
  assumes a: "finite S"
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   132
  shows "supp S = atom ` S"
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   133
proof
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   134
  show "supp S \<subseteq> ((atom::'a::at \<Rightarrow> atom) ` S)" 
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   135
    apply(rule supp_is_subset)
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   136
    apply(simp add: supports_def)
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   137
    apply(rule allI)+
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   138
    apply(rule impI)
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   139
    apply(rule swap_fresh_fresh)
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   140
    apply(simp add: fresh_def)
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   141
    apply(simp add: atom_image_supp)
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   142
    apply(subst supp_finite_atom_set)
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   143
    apply(rule finite_imageI)
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   144
    apply(simp add: a)
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   145
    apply(simp)
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   146
    apply(simp add: fresh_def)
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   147
    apply(simp add: atom_image_supp)
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   148
    apply(subst supp_finite_atom_set)
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   149
    apply(rule finite_imageI)
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   150
    apply(simp add: a)
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   151
    apply(simp)
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   152
    apply(rule finite_imageI)
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   153
    apply(simp add: a)
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   154
    done
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   155
next 
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   156
  have "supp ((atom::'a::at \<Rightarrow> atom) ` S) \<subseteq> supp ((op `) (atom::'a::at \<Rightarrow> atom)) \<union> supp S"
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   157
    by (simp add: supp_fun_app)
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   158
  moreover
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   159
  have "supp ((op `) (atom::'a::at \<Rightarrow> atom)) = {}"
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   160
    apply(rule supp_fun_eqvt)
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   161
    apply(perm_simp)
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   162
    apply(simp)
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   163
    done
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   164
  moreover 
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   165
  have "supp ((atom::'a::at \<Rightarrow> atom) ` S) = ((atom::'a::at \<Rightarrow> atom) ` S)" 
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   166
    apply(subst supp_finite_atom_set)
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   167
    apply(rule finite_imageI)
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   168
    apply(simp add: a)
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   169
    apply(simp)
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   170
    done
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   171
  ultimately 
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   172
  show "((atom::'a::at \<Rightarrow> atom) ` S) \<subseteq> supp S" by simp
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   173
qed
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   174
  
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   175
1933
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   176
lemma supp_at_insert:
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   177
  fixes S::"('a::at) set"
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   178
  assumes a: "finite S"
9eab1dfc14d2 moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   179
  shows "supp (insert a S) = supp a \<union> supp S"
1941
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   180
  using a by (simp add: supp_finite_at_set supp_at_base)
Christian Urban <urbanc@in.tum.de>
parents: 1933
diff changeset
   181
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   182
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   183
section {* A swapping operation for concrete atoms *}
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   184
  
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   185
definition
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   186
  flip :: "'a::at_base \<Rightarrow> 'a \<Rightarrow> perm" ("'(_ \<leftrightarrow> _')")
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   187
where
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   188
  "(a \<leftrightarrow> b) = (atom a \<rightleftharpoons> atom b)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   189
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   190
lemma flip_self [simp]: "(a \<leftrightarrow> a) = 0"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   191
  unfolding flip_def by (rule swap_self)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   192
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   193
lemma flip_commute: "(a \<leftrightarrow> b) = (b \<leftrightarrow> a)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   194
  unfolding flip_def by (rule swap_commute)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   195
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   196
lemma minus_flip [simp]: "- (a \<leftrightarrow> b) = (a \<leftrightarrow> b)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   197
  unfolding flip_def by (rule minus_swap)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   198
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   199
lemma add_flip_cancel: "(a \<leftrightarrow> b) + (a \<leftrightarrow> b) = 0"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   200
  unfolding flip_def by (rule swap_cancel)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   201
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   202
lemma permute_flip_cancel [simp]: "(a \<leftrightarrow> b) \<bullet> (a \<leftrightarrow> b) \<bullet> x = x"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   203
  unfolding permute_plus [symmetric] add_flip_cancel by simp
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   204
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   205
lemma permute_flip_cancel2 [simp]: "(a \<leftrightarrow> b) \<bullet> (b \<leftrightarrow> a) \<bullet> x = x"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   206
  by (simp add: flip_commute)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   207
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   208
lemma flip_eqvt: 
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   209
  fixes a b c::"'a::at_base"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   210
  shows "p \<bullet> (a \<leftrightarrow> b) = (p \<bullet> a \<leftrightarrow> p \<bullet> b)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   211
  unfolding flip_def
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   212
  by (simp add: swap_eqvt atom_eqvt)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   213
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   214
lemma flip_at_base_simps [simp]:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   215
  shows "sort_of (atom a) = sort_of (atom b) \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> a = b"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   216
  and   "sort_of (atom a) = sort_of (atom b) \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> b = a"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   217
  and   "\<lbrakk>a \<noteq> c; b \<noteq> c\<rbrakk> \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> c = c"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   218
  and   "sort_of (atom a) \<noteq> sort_of (atom b) \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> x = x"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   219
  unfolding flip_def
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   220
  unfolding atom_eq_iff [symmetric]
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   221
  unfolding atom_eqvt [symmetric]
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   222
  by simp_all
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   223
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   224
text {* the following two lemmas do not hold for at_base, 
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   225
  only for single sort atoms from at *}
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   226
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   227
lemma permute_flip_at:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   228
  fixes a b c::"'a::at"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   229
  shows "(a \<leftrightarrow> b) \<bullet> c = (if c = a then b else if c = b then a else c)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   230
  unfolding flip_def
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   231
  apply (rule atom_eq_iff [THEN iffD1])
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   232
  apply (subst atom_eqvt [symmetric])
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   233
  apply (simp add: swap_atom)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   234
  done
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   235
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   236
lemma flip_at_simps [simp]:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   237
  fixes a b::"'a::at"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   238
  shows "(a \<leftrightarrow> b) \<bullet> a = b" 
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   239
  and   "(a \<leftrightarrow> b) \<bullet> b = a"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   240
  unfolding permute_flip_at by simp_all
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   241
1499
21dda372fb11 simplified strong induction proof by using flip
Christian Urban <urbanc@in.tum.de>
parents: 1363
diff changeset
   242
lemma flip_fresh_fresh:
21dda372fb11 simplified strong induction proof by using flip
Christian Urban <urbanc@in.tum.de>
parents: 1363
diff changeset
   243
  fixes a b::"'a::at_base"
21dda372fb11 simplified strong induction proof by using flip
Christian Urban <urbanc@in.tum.de>
parents: 1363
diff changeset
   244
  assumes "atom a \<sharp> x" "atom b \<sharp> x"
21dda372fb11 simplified strong induction proof by using flip
Christian Urban <urbanc@in.tum.de>
parents: 1363
diff changeset
   245
  shows "(a \<leftrightarrow> b) \<bullet> x = x"
21dda372fb11 simplified strong induction proof by using flip
Christian Urban <urbanc@in.tum.de>
parents: 1363
diff changeset
   246
using assms
21dda372fb11 simplified strong induction proof by using flip
Christian Urban <urbanc@in.tum.de>
parents: 1363
diff changeset
   247
by (simp add: flip_def swap_fresh_fresh)
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   248
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   249
subsection {* Syntax for coercing at-elements to the atom-type *}
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   250
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   251
syntax
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   252
  "_atom_constrain" :: "logic \<Rightarrow> type \<Rightarrow> logic" ("_:::_" [4, 0] 3)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   253
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   254
translations
1569
1694f32b480a some tuning
Christian Urban <urbanc@in.tum.de>
parents: 1499
diff changeset
   255
  "_atom_constrain a t" => "CONST atom (_constrain a t)"
1694f32b480a some tuning
Christian Urban <urbanc@in.tum.de>
parents: 1499
diff changeset
   256
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   257
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   258
subsection {* A lemma for proving instances of class @{text at}. *}
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   259
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   260
setup {* Sign.add_const_constraint (@{const_name "permute"}, NONE) *}
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   261
setup {* Sign.add_const_constraint (@{const_name "atom"}, NONE) *}
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   262
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   263
text {*
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   264
  New atom types are defined as subtypes of @{typ atom}.
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   265
*}
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   266
1079
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   267
lemma exists_eq_simple_sort: 
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   268
  shows "\<exists>a. a \<in> {a. sort_of a = s}"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   269
  by (rule_tac x="Atom s 0" in exI, simp)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   270
1079
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   271
lemma exists_eq_sort: 
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   272
  shows "\<exists>a. a \<in> {a. sort_of a \<in> range sort_fun}"
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   273
  by (rule_tac x="Atom (sort_fun x) y" in exI, simp)
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   274
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   275
lemma at_base_class:
1079
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   276
  fixes sort_fun :: "'b \<Rightarrow>atom_sort"
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   277
  fixes Rep :: "'a \<Rightarrow> atom" and Abs :: "atom \<Rightarrow> 'a"
1079
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   278
  assumes type: "type_definition Rep Abs {a. sort_of a \<in> range sort_fun}"
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   279
  assumes atom_def: "\<And>a. atom a = Rep a"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   280
  assumes permute_def: "\<And>p a. p \<bullet> a = Abs (p \<bullet> Rep a)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   281
  shows "OFCLASS('a, at_base_class)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   282
proof
1079
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   283
  interpret type_definition Rep Abs "{a. sort_of a \<in> range sort_fun}" by (rule type)
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   284
  have sort_of_Rep: "\<And>a. sort_of (Rep a) \<in> range sort_fun" using Rep by simp
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   285
  fix a b :: 'a and p p1 p2 :: perm
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   286
  show "0 \<bullet> a = a"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   287
    unfolding permute_def by (simp add: Rep_inverse)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   288
  show "(p1 + p2) \<bullet> a = p1 \<bullet> p2 \<bullet> a"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   289
    unfolding permute_def by (simp add: Abs_inverse sort_of_Rep)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   290
  show "atom a = atom b \<longleftrightarrow> a = b"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   291
    unfolding atom_def by (simp add: Rep_inject)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   292
  show "p \<bullet> atom a = atom (p \<bullet> a)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   293
    unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   294
qed
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   295
1079
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   296
(*
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   297
lemma at_class:
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   298
  fixes s :: atom_sort
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   299
  fixes Rep :: "'a \<Rightarrow> atom" and Abs :: "atom \<Rightarrow> 'a"
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   300
  assumes type: "type_definition Rep Abs {a. sort_of a \<in> range (\<lambda>x::unit. s)}"
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   301
  assumes atom_def: "\<And>a. atom a = Rep a"
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   302
  assumes permute_def: "\<And>p a. p \<bullet> a = Abs (p \<bullet> Rep a)"
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   303
  shows "OFCLASS('a, at_class)"
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   304
proof
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   305
  interpret type_definition Rep Abs "{a. sort_of a \<in> range (\<lambda>x::unit. s)}" by (rule type)
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   306
  have sort_of_Rep: "\<And>a. sort_of (Rep a) = s" using Rep by (simp add: image_def)
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   307
  fix a b :: 'a and p p1 p2 :: perm
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   308
  show "0 \<bullet> a = a"
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   309
    unfolding permute_def by (simp add: Rep_inverse)
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   310
  show "(p1 + p2) \<bullet> a = p1 \<bullet> p2 \<bullet> a"
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   311
    unfolding permute_def by (simp add: Abs_inverse sort_of_Rep)
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   312
  show "sort_of (atom a) = sort_of (atom b)"
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   313
    unfolding atom_def by (simp add: sort_of_Rep)
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   314
  show "atom a = atom b \<longleftrightarrow> a = b"
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   315
    unfolding atom_def by (simp add: Rep_inject)
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   316
  show "p \<bullet> atom a = atom (p \<bullet> a)"
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   317
    unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep)
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   318
qed
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   319
*)
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   320
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   321
lemma at_class:
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   322
  fixes s :: atom_sort
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   323
  fixes Rep :: "'a \<Rightarrow> atom" and Abs :: "atom \<Rightarrow> 'a"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   324
  assumes type: "type_definition Rep Abs {a. sort_of a = s}"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   325
  assumes atom_def: "\<And>a. atom a = Rep a"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   326
  assumes permute_def: "\<And>p a. p \<bullet> a = Abs (p \<bullet> Rep a)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   327
  shows "OFCLASS('a, at_class)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   328
proof
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   329
  interpret type_definition Rep Abs "{a. sort_of a = s}" by (rule type)
1079
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   330
  have sort_of_Rep: "\<And>a. sort_of (Rep a) = s" using Rep by (simp add: image_def)
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   331
  fix a b :: 'a and p p1 p2 :: perm
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   332
  show "0 \<bullet> a = a"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   333
    unfolding permute_def by (simp add: Rep_inverse)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   334
  show "(p1 + p2) \<bullet> a = p1 \<bullet> p2 \<bullet> a"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   335
    unfolding permute_def by (simp add: Abs_inverse sort_of_Rep)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   336
  show "sort_of (atom a) = sort_of (atom b)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   337
    unfolding atom_def by (simp add: sort_of_Rep)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   338
  show "atom a = atom b \<longleftrightarrow> a = b"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   339
    unfolding atom_def by (simp add: Rep_inject)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   340
  show "p \<bullet> atom a = atom (p \<bullet> a)"
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   341
    unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep)
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   342
qed
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   343
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   344
setup {* Sign.add_const_constraint
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   345
  (@{const_name "permute"}, SOME @{typ "perm \<Rightarrow> 'a::pt \<Rightarrow> 'a"}) *}
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   346
setup {* Sign.add_const_constraint
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   347
  (@{const_name "atom"}, SOME @{typ "'a::at_base \<Rightarrow> atom"}) *}
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   348
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   349
section {* Automation for creating concrete atom types *}
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   350
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   351
text {* at the moment only single-sort concrete atoms are supported *}
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   352
1079
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   353
use "nominal_atoms.ML"
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   354
c70e7545b738 updated to latest Nominal2
Christian Urban <urbanc@in.tum.de>
parents: 1062
diff changeset
   355
1062
dfea9e739231 rollback of the test
Christian Urban <urbanc@in.tum.de>
parents: 1061
diff changeset
   356
end