section ‹Enumerations›
theory Enumeration_ZF imports NatOrder_ZF FiniteSeq_ZF FinOrd_ZF
begin
text‹Suppose $r$ is a linear order on a set $A$ that has $n$ elements,
where $n \in\mathbb{N}$ . In the ‹FinOrd_ZF›
theory we prove a theorem stating that there is a unique
order isomorphism between $n = \{0,1,..,n-1\}$ (with natural order) and
$A$. Another way of stating that is that there is
a unique way of counting the elements of $A$ in the order increasing
according to relation $r$. Yet another way of stating the same thing is
that there is a unique sorted list of elements of $A$. We will call this
list the ‹Enumeration› of $A$.›
subsection‹Enumerations: definition and notation›
text‹In this section we introduce the notion of enumeration
and define a proof context (a ''locale'' in Isabelle terms)
that sets up the notation for writing about enumarations.›
text‹We define enumeration as the only order isomorphism beween
a set $A$ and the number of its elements. We are using
the formula $\bigcup \{x\} = x$ to extract the only element
from a singleton. ‹Le› is the (natural) order on natural
numbers, defined is ‹Nat_ZF› theory in the standard
Isabelle library.›
definition
"Enumeration(A,r) ≡ ⋃ ord_iso(|A|,Le,A,r)"
text‹To set up the notation we define a locale ‹enums›. In this
locale we will assume that $r$ is a linear order on some set $X$.
In most applications this set will be just the set of natural numbers.
Standard Isabelle uses $\leq $ to denote the "less or equal" relation
on natural numbers. We will use the ‹\<lsq>› symbol to denote
the relation $r$. Those two symbols usually look the same in the presentation,
but they are different in the source.To shorten the notation the enumeration
‹Enumeration(A,r)› will be denoted as $\sigma (A)$.
Similarly as in the ‹Semigroup› theory we will write
$a\hookleftarrow x$ for the result of appending an element $x$ to
the finite sequence (list) $a$. Finally,
$a\sqcup b$ will denote the concatenation of the lists $a$ and $b$.›
locale enums =
fixes X r
assumes linord: "IsLinOrder(X,r)"
fixes ler (infix "\<lsq>" 70)
defines ler_def[simp]: "x \<lsq> y ≡ ⟨x,y⟩ ∈ r"
fixes σ
defines σ_def [simp]: "σ(A) ≡ Enumeration(A,r)"
fixes append (infix "↩" 72)
defines append_def[simp]: "a ↩ x ≡ Append(a,x)"
fixes concat (infixl "⊔" 69)
defines concat_def[simp]: "a ⊔ b ≡ Concat(a,b)"
subsection‹Properties of enumerations›
text‹In this section we prove basic facts about enumerations.›
text‹A special case of the existence and uniqueess
of the order isomorphism for finite sets
when the first set is a natural number.›
lemma (in enums) ord_iso_nat_fin:
assumes "A ∈ FinPow(X)" and "n ∈ nat" and "A ≈ n"
shows "∃!f. f ∈ ord_iso(n,Le,A,r)"
using assms NatOrder_ZF_1_L2 linord nat_finpow_nat
fin_ord_iso_ex_uniq by simp
text‹An enumeration is an order isomorhism, a bijection, and a list.›
lemma (in enums) enum_props: assumes "A ∈ FinPow(X)"
shows
"σ(A) ∈ ord_iso(|A|,Le, A,r)"
"σ(A) ∈ bij(|A|,A)"
"σ(A) : |A| → A"
proof -
from assms have
"IsLinOrder(nat,Le)" and "|A| ∈ FinPow(nat)" and "A ≈ |A|"
using NatOrder_ZF_1_L2 card_fin_is_nat nat_finpow_nat
by auto
with assms show "σ(A) ∈ ord_iso(|A|,Le, A,r)"
using linord fin_ord_iso_ex_uniq singleton_extract
Enumeration_def by simp
then show "σ(A) ∈ bij(|A|,A)" and "σ(A) : |A| → A"
using ord_iso_def bij_def surj_def
by auto
qed
text‹A corollary from ‹enum_props›. Could have been attached as
another assertion, but this slows down verification of some other proofs.
›
lemma (in enums) enum_fun: assumes "A ∈ FinPow(X)"
shows "σ(A) : |A| → X"
proof -
from assms have "σ(A) : |A| → A" and "A⊆X"
using enum_props FinPow_def by auto
then show "σ(A) : |A| → X" by (rule func1_1_L1B)
qed
text‹If a list is an order isomorphism then it must be the enumeration.
›
lemma (in enums) ord_iso_enum: assumes A1: "A ∈ FinPow(X)" and
A2: "n ∈ nat" and A3: "f ∈ ord_iso(n,Le,A,r)"
shows "f = σ(A)"
proof -
from A3 have "n ≈ A" using ord_iso_def eqpoll_def
by auto
then have "A ≈ n" by (rule eqpoll_sym)
with A1 A2 have "∃!f. f ∈ ord_iso(n,Le,A,r)"
using ord_iso_nat_fin by simp
with assms ‹A ≈ n› show "f = σ(A)"
using enum_props card_card by blast
qed
text‹What is the enumeration of the empty set?›
lemma (in enums) empty_enum: shows "σ(0) = 0"
proof -
have
"0 ∈ FinPow(X)" and "0 ∈ nat" and "0 ∈ ord_iso(0,Le,0,r)"
using empty_in_finpow empty_ord_iso_empty
by auto
then show "σ(0) = 0" using ord_iso_enum
by blast
qed
text‹Adding a new maximum to a set appends it to the enumeration.›
lemma (in enums) enum_append:
assumes A1: "A ∈ FinPow(X)" and A2: "b ∈ X-A" and
A3: "∀a∈A. a\<lsq>b"
shows " σ(A ∪ {b}) = σ(A)↩ b"
proof -
let ?f = "σ(A) ∪ {⟨|A|,b⟩}"
from A1 have "|A| ∈ nat" using card_fin_is_nat
by simp
from A1 A2 have "A ∪ {b} ∈ FinPow(X)"
using singleton_in_finpow union_finpow by simp
moreover from this have "|A ∪ {b}| ∈ nat"
using card_fin_is_nat by simp
moreover have "?f ∈ ord_iso(|A ∪ {b}| , Le, A ∪ {b} ,r)"
proof -
from A1 A2 have
"σ(A) ∈ ord_iso(|A|,Le, A,r)" and
"|A| ∉ |A|" and "b ∉ A"
using enum_props mem_not_refl by auto
moreover from ‹|A| ∈ nat› have
"∀k ∈ |A|. ⟨k, |A|⟩ ∈ Le"
using elem_nat_is_nat by blast
moreover from A3 have "∀a∈A. ⟨a,b⟩ ∈ r" by simp
moreover have "antisym(Le)" and "antisym(r)"
using linord NatOrder_ZF_1_L2 IsLinOrder_def by auto
moreover
from A2 ‹|A| ∈ nat› have
"⟨|A|,|A|⟩ ∈ Le" and "⟨b,b⟩ ∈ r"
using linord NatOrder_ZF_1_L2 IsLinOrder_def
total_is_refl refl_def by auto
hence "⟨|A|,|A|⟩ ∈ Le ⟷ ⟨b,b⟩ ∈ r" by simp
ultimately have "?f ∈ ord_iso(|A| ∪ {|A|} , Le, A ∪ {b} ,r)"
by (rule ord_iso_extend)
with A1 A2 show "?f ∈ ord_iso(|A ∪ {b}| , Le, A ∪ {b} ,r)"
using card_fin_add_one by simp
qed
ultimately have "?f = σ(A ∪ {b})"
using ord_iso_enum by simp
moreover have "σ(A)↩ b = ?f"
proof -
have "σ(A)↩ b = σ(A) ∪ {⟨domain(σ(A)),b⟩}"
using Append_def by simp
moreover from A1 have "domain(σ(A)) = |A|"
using enum_props func1_1_L1 by blast
ultimately show "σ(A)↩ b = ?f" by simp
qed
ultimately show "σ(A ∪ {b}) = σ(A)↩ b" by simp
qed
text‹What is the enumeration of a singleton?›
lemma (in enums) enum_singleton:
assumes A1: "x∈X" shows "σ({x}): 1 → X" and "σ({x})`(0) = x"
proof -
from A1 have
"0 ∈ FinPow(X)" and "x ∈ (X - 0)" and "∀a∈0. a\<lsq>x"
using empty_in_finpow by auto
then have "σ(0 ∪ {x}) = σ(0)↩ x" by (rule enum_append)
with A1 show "σ({x}): 1 → X" and "σ({x})`(0) = x"
using empty_enum empty_append1 by auto
qed
end