Theory Indexing
subsection ‹Indexing›
theory Indexing
imports Main
begin
definition injective :: "nat ⇒ ('k ⇒ nat) ⇒ bool" where
"injective size to_index ≡ ∀ a b.
to_index a = to_index b
∧ to_index a < size
∧ to_index b < size
⟶ a = b"
for size to_index
lemma index_mono:
fixes a b a0 b0 :: nat
assumes a: "a < a0" and b: "b < b0"
shows "a * b0 + b < a0 * b0"
proof -
have "a * b0 + b < (Suc a) * b0"
using b by auto
also have "… ≤ a0 * b0"
using a[THEN Suc_leI, THEN mult_le_mono1] .
finally show ?thesis .
qed
lemma index_eq_iff:
fixes a b c d b0 :: nat
assumes "b < b0" "d < b0" "a * b0 + b = c * b0 + d"
shows "a = c ∧ b = d"
proof (rule conjI)
{ fix a b c d :: nat
assume ac: "a < c" and b: "b < b0"
have "a * b0 + b < (Suc a) * b0"
using b by auto
also have "… ≤ c * b0"
using ac[THEN Suc_leI, THEN mult_le_mono1] .
also have "… ≤ c * b0 + d"
by auto
finally have "a * b0 + b ≠ c * b0 + d"
by auto
} note ac = this
{ assume "a ≠ c"
then consider (le) "a < c" | (ge) "a > c"
by fastforce
hence False proof cases
case le show ?thesis using ac[OF le assms(1)] assms(3) ..
next
case ge show ?thesis using ac[OF ge assms(2)] assms(3)[symmetric] ..
qed
}
then show "a = c"
by auto
with assms(3) show "b = d"
by auto
qed
locale prod_order_def =
order0: ord less_eq0 less0 +
order1: ord less_eq1 less1
for less_eq0 less0 less_eq1 less1
begin
fun less :: "'a × 'b ⇒ 'a × 'b ⇒ bool" where
"less (a,b) (c,d) ⟷ less0 a c ∧ less1 b d"
fun less_eq :: "'a × 'b ⇒ 'a × 'b ⇒ bool" where
"less_eq ab cd ⟷ less ab cd ∨ ab = cd"
end
locale prod_order =
prod_order_def less_eq0 less0 less_eq1 less1 +
order0: order less_eq0 less0 +
order1: order less_eq1 less1
for less_eq0 less0 less_eq1 less1
begin
sublocale order less_eq less
proof qed fastforce+
end
locale option_order =
order0: order less_eq0 less0
for less_eq0 less0
begin
fun less_eq_option :: "'a option ⇒ 'a option ⇒ bool" where
"less_eq_option None _ ⟷ True"
| "less_eq_option (Some _) None ⟷ False"
| "less_eq_option (Some a) (Some b) ⟷ less_eq0 a b"
fun less_option :: "'a option ⇒ 'a option ⇒ bool" where
"less_option ao bo ⟷ less_eq_option ao bo ∧ ao ≠ bo"
sublocale order less_eq_option less_option
apply standard
subgoal for x y by (cases x; cases y) auto
subgoal for x by (cases x) auto
subgoal for x y z by (cases x; cases y; cases z) auto
subgoal for x y by (cases x; cases y) auto
done
end