section ‹Order on natural numbers›
theory NatOrder_ZF imports Nat_ZF_IML Order_ZF
begin
text‹This theory proves that $\leq$ is a linear order on $\mathbb{N}$.
$\leq$ is defined in Isabelle's ‹Nat› theory, and
linear order is defined in ‹Order_ZF› theory.
Contributed by Seo Sanghyeon.›
subsection‹Order on natural numbers›
text‹This is the only section in this theory.›
text‹To prove that $\leq$ is a total order, we use a result on ordinals.›
lemma NatOrder_ZF_1_L1:
assumes "a∈nat" and "b∈nat"
shows "a ≤ b ∨ b ≤ a"
proof -
from assms have I: "Ord(a) ∧ Ord(b)"
using nat_into_Ord by auto
then have "a ∈ b ∨ a = b ∨ b ∈ a"
using Ord_linear by simp
with I have "a < b ∨ a = b ∨ b < a"
using ltI by auto
with I show "a ≤ b ∨ b ≤ a"
using le_iff by auto
qed
text‹$\leq$ is antisymmetric, transitive, total, and linear. Proofs by
rewrite using definitions.›
lemma NatOrder_ZF_1_L2:
shows
"antisym(Le)"
"trans(Le)"
"Le {is total on} nat"
"IsLinOrder(nat,Le)"
proof -
show "antisym(Le)"
using antisym_def Le_def le_anti_sym by auto
moreover show "trans(Le)"
using trans_def Le_def le_trans by blast
moreover show "Le {is total on} nat"
using IsTotal_def Le_def NatOrder_ZF_1_L1 by simp
ultimately show "IsLinOrder(nat,Le)"
using IsLinOrder_def by simp
qed
text‹The order on natural numbers is linear on every natural number.
Recall that each natural number is a subset of the set of
all natural numbers (as well as a member).›
lemma natord_lin_on_each_nat:
assumes A1: "n ∈ nat" shows "IsLinOrder(n,Le)"
proof -
from A1 have "n ⊆ nat" using nat_subset_nat
by simp
then show ?thesis using NatOrder_ZF_1_L2 ord_linear_subset
by blast
qed
end