header {* Miscellaneous Tools for Size-Change Termination *}
theory Misc_Tools
imports Main
begin
subsection {* Searching in lists *}
fun index_of :: "'a list => 'a => nat"
where
"index_of [] c = 0"
| "index_of (x#xs) c = (if x = c then 0 else Suc (index_of xs c))"
lemma index_of_member:
"(x ∈ set l) ==> (l ! index_of l x = x)"
by (induct l) auto
lemma index_of_length:
"(x ∈ set l) = (index_of l x < length l)"
by (induct l) auto
subsection {* Some reasoning tools *}
lemma three_cases:
assumes "a1 ==> thesis"
assumes "a2 ==> thesis"
assumes "a3 ==> thesis"
assumes "!!R. [|a1 ==> R; a2 ==> R; a3 ==> R|] ==> R"
shows "thesis"
using assms
by auto
subsection {* Sequences *}
types
'a sequence = "nat => 'a"
subsubsection {* Increasing sequences *}
definition
increasing :: "(nat => nat) => bool" where
"increasing s = (∀i j. i < j --> s i < s j)"
lemma increasing_strict:
assumes "increasing s"
assumes "i < j"
shows "s i < s j"
using assms
unfolding increasing_def by simp
lemma increasing_weak:
assumes "increasing s"
assumes "i ≤ j"
shows "s i ≤ s j"
using assms increasing_strict[of s i j]
by (cases "i < j") auto
lemma increasing_inc:
assumes "increasing s"
shows "n ≤ s n"
proof (induct n)
case 0 then show ?case by simp
next
case (Suc n)
with increasing_strict [OF `increasing s`, of n "Suc n"]
show ?case by auto
qed
lemma increasing_bij:
assumes [simp]: "increasing s"
shows "(s i < s j) = (i < j)"
proof
assume "s i < s j"
show "i < j"
proof (rule classical)
assume "¬ ?thesis"
hence "j ≤ i" by arith
with increasing_weak have "s j ≤ s i" by simp
with `s i < s j` show ?thesis by simp
qed
qed (simp add:increasing_strict)
subsubsection {* Sections induced by an increasing sequence *}
abbreviation
"section s i == {s i ..< s (Suc i)}"
definition
"section_of s n = (LEAST i. n < s (Suc i))"
lemma section_help:
assumes "increasing s"
shows "∃i. n < s (Suc i)"
proof -
have "n ≤ s n"
using `increasing s` by (rule increasing_inc)
also have "… < s (Suc n)"
using `increasing s` increasing_strict by simp
finally show ?thesis ..
qed
lemma section_of2:
assumes "increasing s"
shows "n < s (Suc (section_of s n))"
unfolding section_of_def
by (rule LeastI_ex) (rule section_help [OF `increasing s`])
lemma section_of1:
assumes [simp, intro]: "increasing s"
assumes "s i ≤ n"
shows "s (section_of s n) ≤ n"
proof (rule classical)
let ?m = "section_of s n"
assume "¬ ?thesis"
hence a: "n < s ?m" by simp
have nonzero: "?m ≠ 0"
proof
assume "?m = 0"
from increasing_weak have "s 0 ≤ s i" by simp
also note `… ≤ n`
finally show False using `?m = 0` `n < s ?m` by simp
qed
with a have "n < s (Suc (?m - 1))" by simp
with Least_le have "?m ≤ ?m - 1"
unfolding section_of_def .
with nonzero show ?thesis by simp
qed
lemma section_of_known:
assumes [simp]: "increasing s"
assumes in_sect: "k ∈ section s i"
shows "section_of s k = i" (is "?s = i")
proof (rule classical)
assume "¬ ?thesis"
hence "?s < i ∨ ?s > i" by arith
thus ?thesis
proof
assume "?s < i"
hence "Suc ?s ≤ i" by simp
with increasing_weak have "s (Suc ?s) ≤ s i" by simp
moreover have "k < s (Suc ?s)" using section_of2 by simp
moreover from in_sect have "s i ≤ k" by simp
ultimately show ?thesis by simp
next
assume "i < ?s" hence "Suc i ≤ ?s" by simp
with increasing_weak have "s (Suc i) ≤ s ?s" by simp
moreover
from in_sect have "s i ≤ k" by simp
with section_of1 have "s ?s ≤ k" by simp
moreover from in_sect have "k < s (Suc i)" by simp
ultimately show ?thesis by simp
qed
qed
lemma in_section_of:
assumes "increasing s"
assumes "s i ≤ k"
shows "k ∈ section s (section_of s k)"
using assms
by (auto intro:section_of1 section_of2)
end