1062
+ − 1
(* Title: Nominal2_Atoms
+ − 2
Authors: Brian Huffman, Christian Urban
+ − 3
+ − 4
Definitions for concrete atom types.
+ − 5
*)
+ − 6
theory Nominal2_Atoms
+ − 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>
diff
changeset
+ − 8
Nominal2_Eqvt
1079
+ − 9
uses ("nominal_atoms.ML")
1062
+ − 10
begin
+ − 11
+ − 12
section {* Concrete atom types *}
+ − 13
+ − 14
text {*
+ − 15
Class @{text at_base} allows types containing multiple sorts of atoms.
+ − 16
Class @{text at} only allows types with a single sort.
+ − 17
*}
+ − 18
+ − 19
class at_base = pt +
+ − 20
fixes atom :: "'a \<Rightarrow> atom"
+ − 21
assumes atom_eq_iff [simp]: "atom a = atom b \<longleftrightarrow> a = b"
+ − 22
assumes atom_eqvt: "p \<bullet> (atom a) = atom (p \<bullet> a)"
+ − 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>
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>
diff
changeset
+ − 25
1062
+ − 26
class at = at_base +
+ − 27
assumes sort_of_atom_eq [simp]: "sort_of (atom a) = sort_of (atom b)"
+ − 28
+ − 29
lemma supp_at_base:
+ − 30
fixes a::"'a::at_base"
+ − 31
shows "supp a = {atom a}"
+ − 32
by (simp add: supp_atom [symmetric] supp_def atom_eqvt)
+ − 33
1079
+ − 34
lemma fresh_at_base:
1062
+ − 35
shows "a \<sharp> b \<longleftrightarrow> a \<noteq> atom b"
+ − 36
unfolding fresh_def by (simp add: supp_at_base)
+ − 37
+ − 38
instance at_base < fs
+ − 39
proof qed (simp add: supp_at_base)
+ − 40
+ − 41
lemma at_base_infinite [simp]:
+ − 42
shows "infinite (UNIV :: 'a::at_base set)" (is "infinite ?U")
+ − 43
proof
+ − 44
obtain a :: 'a where "True" by auto
+ − 45
assume "finite ?U"
+ − 46
hence "finite (atom ` ?U)"
+ − 47
by (rule finite_imageI)
+ − 48
then obtain b where b: "b \<notin> atom ` ?U" "sort_of b = sort_of (atom a)"
+ − 49
by (rule obtain_atom)
+ − 50
from b(2) have "b = atom ((atom a \<rightleftharpoons> b) \<bullet> a)"
+ − 51
unfolding atom_eqvt [symmetric]
+ − 52
by (simp add: swap_atom)
+ − 53
hence "b \<in> atom ` ?U" by simp
+ − 54
with b(1) show "False" by simp
+ − 55
qed
+ − 56
+ − 57
lemma swap_at_base_simps [simp]:
+ − 58
fixes x y::"'a::at_base"
+ − 59
shows "sort_of (atom x) = sort_of (atom y) \<Longrightarrow> (atom x \<rightleftharpoons> atom y) \<bullet> x = y"
+ − 60
and "sort_of (atom x) = sort_of (atom y) \<Longrightarrow> (atom x \<rightleftharpoons> atom y) \<bullet> y = x"
+ − 61
and "atom x \<noteq> a \<Longrightarrow> atom x \<noteq> b \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> x = x"
+ − 62
unfolding atom_eq_iff [symmetric]
+ − 63
unfolding atom_eqvt [symmetric]
+ − 64
by simp_all
+ − 65
+ − 66
lemma obtain_at_base:
+ − 67
assumes X: "finite X"
+ − 68
obtains a::"'a::at_base" where "atom a \<notin> X"
+ − 69
proof -
+ − 70
have "inj (atom :: 'a \<Rightarrow> atom)"
+ − 71
by (simp add: inj_on_def)
+ − 72
with X have "finite (atom -` X :: 'a set)"
+ − 73
by (rule finite_vimageI)
+ − 74
with at_base_infinite have "atom -` X \<noteq> (UNIV :: 'a set)"
+ − 75
by auto
+ − 76
then obtain a :: 'a where "atom a \<notin> X"
+ − 77
by auto
+ − 78
thus ?thesis ..
+ − 79
qed
+ − 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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
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>
diff
changeset
+ − 128
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>
diff
changeset
+ − 129
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>
diff
changeset
+ − 130
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>
diff
changeset
+ − 131
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>
diff
changeset
+ − 132
shows "supp (insert a S) = supp a \<union> supp 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>
diff
changeset
+ − 133
using 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>
diff
changeset
+ − 134
apply (simp add: 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>
diff
changeset
+ − 135
apply (simp add: supp_at_base)
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>
diff
changeset
+ − 136
done
1062
+ − 137
+ − 138
section {* A swapping operation for concrete atoms *}
+ − 139
+ − 140
definition
+ − 141
flip :: "'a::at_base \<Rightarrow> 'a \<Rightarrow> perm" ("'(_ \<leftrightarrow> _')")
+ − 142
where
+ − 143
"(a \<leftrightarrow> b) = (atom a \<rightleftharpoons> atom b)"
+ − 144
+ − 145
lemma flip_self [simp]: "(a \<leftrightarrow> a) = 0"
+ − 146
unfolding flip_def by (rule swap_self)
+ − 147
+ − 148
lemma flip_commute: "(a \<leftrightarrow> b) = (b \<leftrightarrow> a)"
+ − 149
unfolding flip_def by (rule swap_commute)
+ − 150
+ − 151
lemma minus_flip [simp]: "- (a \<leftrightarrow> b) = (a \<leftrightarrow> b)"
+ − 152
unfolding flip_def by (rule minus_swap)
+ − 153
+ − 154
lemma add_flip_cancel: "(a \<leftrightarrow> b) + (a \<leftrightarrow> b) = 0"
+ − 155
unfolding flip_def by (rule swap_cancel)
+ − 156
+ − 157
lemma permute_flip_cancel [simp]: "(a \<leftrightarrow> b) \<bullet> (a \<leftrightarrow> b) \<bullet> x = x"
+ − 158
unfolding permute_plus [symmetric] add_flip_cancel by simp
+ − 159
+ − 160
lemma permute_flip_cancel2 [simp]: "(a \<leftrightarrow> b) \<bullet> (b \<leftrightarrow> a) \<bullet> x = x"
+ − 161
by (simp add: flip_commute)
+ − 162
+ − 163
lemma flip_eqvt:
+ − 164
fixes a b c::"'a::at_base"
+ − 165
shows "p \<bullet> (a \<leftrightarrow> b) = (p \<bullet> a \<leftrightarrow> p \<bullet> b)"
+ − 166
unfolding flip_def
+ − 167
by (simp add: swap_eqvt atom_eqvt)
+ − 168
+ − 169
lemma flip_at_base_simps [simp]:
+ − 170
shows "sort_of (atom a) = sort_of (atom b) \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> a = b"
+ − 171
and "sort_of (atom a) = sort_of (atom b) \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> b = a"
+ − 172
and "\<lbrakk>a \<noteq> c; b \<noteq> c\<rbrakk> \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> c = c"
+ − 173
and "sort_of (atom a) \<noteq> sort_of (atom b) \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> x = x"
+ − 174
unfolding flip_def
+ − 175
unfolding atom_eq_iff [symmetric]
+ − 176
unfolding atom_eqvt [symmetric]
+ − 177
by simp_all
+ − 178
+ − 179
text {* the following two lemmas do not hold for at_base,
+ − 180
only for single sort atoms from at *}
+ − 181
+ − 182
lemma permute_flip_at:
+ − 183
fixes a b c::"'a::at"
+ − 184
shows "(a \<leftrightarrow> b) \<bullet> c = (if c = a then b else if c = b then a else c)"
+ − 185
unfolding flip_def
+ − 186
apply (rule atom_eq_iff [THEN iffD1])
+ − 187
apply (subst atom_eqvt [symmetric])
+ − 188
apply (simp add: swap_atom)
+ − 189
done
+ − 190
+ − 191
lemma flip_at_simps [simp]:
+ − 192
fixes a b::"'a::at"
+ − 193
shows "(a \<leftrightarrow> b) \<bullet> a = b"
+ − 194
and "(a \<leftrightarrow> b) \<bullet> b = a"
+ − 195
unfolding permute_flip_at by simp_all
+ − 196
1499
+ − 197
lemma flip_fresh_fresh:
+ − 198
fixes a b::"'a::at_base"
+ − 199
assumes "atom a \<sharp> x" "atom b \<sharp> x"
+ − 200
shows "(a \<leftrightarrow> b) \<bullet> x = x"
+ − 201
using assms
+ − 202
by (simp add: flip_def swap_fresh_fresh)
1062
+ − 203
+ − 204
subsection {* Syntax for coercing at-elements to the atom-type *}
+ − 205
+ − 206
syntax
+ − 207
"_atom_constrain" :: "logic \<Rightarrow> type \<Rightarrow> logic" ("_:::_" [4, 0] 3)
+ − 208
+ − 209
translations
1569
+ − 210
"_atom_constrain a t" => "CONST atom (_constrain a t)"
+ − 211
1062
+ − 212
+ − 213
subsection {* A lemma for proving instances of class @{text at}. *}
+ − 214
+ − 215
setup {* Sign.add_const_constraint (@{const_name "permute"}, NONE) *}
+ − 216
setup {* Sign.add_const_constraint (@{const_name "atom"}, NONE) *}
+ − 217
+ − 218
text {*
+ − 219
New atom types are defined as subtypes of @{typ atom}.
+ − 220
*}
+ − 221
1079
+ − 222
lemma exists_eq_simple_sort:
1062
+ − 223
shows "\<exists>a. a \<in> {a. sort_of a = s}"
+ − 224
by (rule_tac x="Atom s 0" in exI, simp)
+ − 225
1079
+ − 226
lemma exists_eq_sort:
+ − 227
shows "\<exists>a. a \<in> {a. sort_of a \<in> range sort_fun}"
+ − 228
by (rule_tac x="Atom (sort_fun x) y" in exI, simp)
+ − 229
1062
+ − 230
lemma at_base_class:
1079
+ − 231
fixes sort_fun :: "'b \<Rightarrow>atom_sort"
1062
+ − 232
fixes Rep :: "'a \<Rightarrow> atom" and Abs :: "atom \<Rightarrow> 'a"
1079
+ − 233
assumes type: "type_definition Rep Abs {a. sort_of a \<in> range sort_fun}"
1062
+ − 234
assumes atom_def: "\<And>a. atom a = Rep a"
+ − 235
assumes permute_def: "\<And>p a. p \<bullet> a = Abs (p \<bullet> Rep a)"
+ − 236
shows "OFCLASS('a, at_base_class)"
+ − 237
proof
1079
+ − 238
interpret type_definition Rep Abs "{a. sort_of a \<in> range sort_fun}" by (rule type)
+ − 239
have sort_of_Rep: "\<And>a. sort_of (Rep a) \<in> range sort_fun" using Rep by simp
1062
+ − 240
fix a b :: 'a and p p1 p2 :: perm
+ − 241
show "0 \<bullet> a = a"
+ − 242
unfolding permute_def by (simp add: Rep_inverse)
+ − 243
show "(p1 + p2) \<bullet> a = p1 \<bullet> p2 \<bullet> a"
+ − 244
unfolding permute_def by (simp add: Abs_inverse sort_of_Rep)
+ − 245
show "atom a = atom b \<longleftrightarrow> a = b"
+ − 246
unfolding atom_def by (simp add: Rep_inject)
+ − 247
show "p \<bullet> atom a = atom (p \<bullet> a)"
+ − 248
unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep)
+ − 249
qed
+ − 250
1079
+ − 251
(*
+ − 252
lemma at_class:
+ − 253
fixes s :: atom_sort
+ − 254
fixes Rep :: "'a \<Rightarrow> atom" and Abs :: "atom \<Rightarrow> 'a"
+ − 255
assumes type: "type_definition Rep Abs {a. sort_of a \<in> range (\<lambda>x::unit. s)}"
+ − 256
assumes atom_def: "\<And>a. atom a = Rep a"
+ − 257
assumes permute_def: "\<And>p a. p \<bullet> a = Abs (p \<bullet> Rep a)"
+ − 258
shows "OFCLASS('a, at_class)"
+ − 259
proof
+ − 260
interpret type_definition Rep Abs "{a. sort_of a \<in> range (\<lambda>x::unit. s)}" by (rule type)
+ − 261
have sort_of_Rep: "\<And>a. sort_of (Rep a) = s" using Rep by (simp add: image_def)
+ − 262
fix a b :: 'a and p p1 p2 :: perm
+ − 263
show "0 \<bullet> a = a"
+ − 264
unfolding permute_def by (simp add: Rep_inverse)
+ − 265
show "(p1 + p2) \<bullet> a = p1 \<bullet> p2 \<bullet> a"
+ − 266
unfolding permute_def by (simp add: Abs_inverse sort_of_Rep)
+ − 267
show "sort_of (atom a) = sort_of (atom b)"
+ − 268
unfolding atom_def by (simp add: sort_of_Rep)
+ − 269
show "atom a = atom b \<longleftrightarrow> a = b"
+ − 270
unfolding atom_def by (simp add: Rep_inject)
+ − 271
show "p \<bullet> atom a = atom (p \<bullet> a)"
+ − 272
unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep)
+ − 273
qed
+ − 274
*)
+ − 275
1062
+ − 276
lemma at_class:
+ − 277
fixes s :: atom_sort
+ − 278
fixes Rep :: "'a \<Rightarrow> atom" and Abs :: "atom \<Rightarrow> 'a"
+ − 279
assumes type: "type_definition Rep Abs {a. sort_of a = s}"
+ − 280
assumes atom_def: "\<And>a. atom a = Rep a"
+ − 281
assumes permute_def: "\<And>p a. p \<bullet> a = Abs (p \<bullet> Rep a)"
+ − 282
shows "OFCLASS('a, at_class)"
+ − 283
proof
+ − 284
interpret type_definition Rep Abs "{a. sort_of a = s}" by (rule type)
1079
+ − 285
have sort_of_Rep: "\<And>a. sort_of (Rep a) = s" using Rep by (simp add: image_def)
1062
+ − 286
fix a b :: 'a and p p1 p2 :: perm
+ − 287
show "0 \<bullet> a = a"
+ − 288
unfolding permute_def by (simp add: Rep_inverse)
+ − 289
show "(p1 + p2) \<bullet> a = p1 \<bullet> p2 \<bullet> a"
+ − 290
unfolding permute_def by (simp add: Abs_inverse sort_of_Rep)
+ − 291
show "sort_of (atom a) = sort_of (atom b)"
+ − 292
unfolding atom_def by (simp add: sort_of_Rep)
+ − 293
show "atom a = atom b \<longleftrightarrow> a = b"
+ − 294
unfolding atom_def by (simp add: Rep_inject)
+ − 295
show "p \<bullet> atom a = atom (p \<bullet> a)"
+ − 296
unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep)
+ − 297
qed
+ − 298
+ − 299
setup {* Sign.add_const_constraint
+ − 300
(@{const_name "permute"}, SOME @{typ "perm \<Rightarrow> 'a::pt \<Rightarrow> 'a"}) *}
+ − 301
setup {* Sign.add_const_constraint
+ − 302
(@{const_name "atom"}, SOME @{typ "'a::at_base \<Rightarrow> atom"}) *}
+ − 303
+ − 304
section {* Automation for creating concrete atom types *}
+ − 305
+ − 306
text {* at the moment only single-sort concrete atoms are supported *}
+ − 307
1079
+ − 308
use "nominal_atoms.ML"
+ − 309
+ − 310
1062
+ − 311
end