Theory ListSpec
section ‹\isaheader{Specification of Sequences}›
theory ListSpec
imports ICF_Spec_Base
begin
subsection "Definition"
locale list =
fixes α :: "'s ⇒ 'x list"
fixes invar :: "'s ⇒ bool"
locale list_no_invar = list +
assumes invar[simp, intro!]: "⋀l. invar l"
subsection "Functions"
locale list_empty = list +
constrains α :: "'s ⇒ 'x list"
fixes empty :: "unit ⇒ 's"
assumes empty_correct:
"α (empty ()) = []"
"invar (empty ())"
locale list_isEmpty = list +
constrains α :: "'s ⇒ 'x list"
fixes isEmpty :: "'s ⇒ bool"
assumes isEmpty_correct:
"invar s ⟹ isEmpty s ⟷ α s = []"
locale poly_list_iteratei = list +
constrains α :: "'s ⇒ 'x list"
begin
definition iteratei where
iteratei_correct[code_unfold]: "iteratei s ≡ foldli (α s)"
definition iterate where
iterate_correct[code_unfold]: "iterate s ≡ foldli (α s) (λ_. True)"
end
locale poly_list_rev_iteratei = list +
constrains α :: "'s ⇒ 'x list"
begin
definition rev_iteratei where
rev_iteratei_correct[code_unfold]: "rev_iteratei s ≡ foldri (α s)"
definition rev_iterate where
rev_iterate_correct[code_unfold]: "rev_iterate s ≡ foldri (α s) (λ_. True)"
end
locale list_size = list +
constrains α :: "'s ⇒ 'x list"
fixes size :: "'s ⇒ nat"
assumes size_correct:
"invar s ⟹ size s = length (α s)"
locale list_appendl = list +
constrains α :: "'s ⇒ 'x list"
fixes appendl :: "'x ⇒ 's ⇒ 's"
assumes appendl_correct:
"invar s ⟹ α (appendl x s) = x#α s"
"invar s ⟹ invar (appendl x s)"
begin
abbreviation (input) "push ≡ appendl"
lemmas push_correct = appendl_correct
end
locale list_removel = list +
constrains α :: "'s ⇒ 'x list"
fixes removel :: "'s ⇒ ('x × 's)"
assumes removel_correct:
"⟦invar s; α s ≠ []⟧ ⟹ fst (removel s) = hd (α s)"
"⟦invar s; α s ≠ []⟧ ⟹ α (snd (removel s)) = tl (α s)"
"⟦invar s; α s ≠ []⟧ ⟹ invar (snd (removel s))"
begin
lemma removelE:
assumes I: "invar s" "α s ≠ []"
obtains s' where "removel s = (hd (α s), s')" "invar s'" "α s' = tl (α s)"
proof -
from removel_correct(1,2,3)[OF I] have
C: "fst (removel s) = hd (α s)"
"α (snd (removel s)) = tl (α s)"
"invar (snd (removel s))" .
from that[of "snd (removel s)", OF _ C(3,2), folded C(1)] show thesis
by simp
qed
text ‹The following shortcut notations are not meant for generating efficient code,
but solely to simplify reasoning›
abbreviation (input) "pop ≡ removel"
lemmas pop_correct = removel_correct
abbreviation (input) "dequeue ≡ removel"
lemmas dequeue_correct = removel_correct
end
locale list_leftmost = list +
constrains α :: "'s ⇒ 'x list"
fixes leftmost :: "'s ⇒ 'x"
assumes leftmost_correct:
"⟦invar s; α s ≠ []⟧ ⟹ leftmost s = hd (α s)"
begin
abbreviation (input) top where "top ≡ leftmost"
lemmas top_correct = leftmost_correct
end
locale list_appendr = list +
constrains α :: "'s ⇒ 'x list"
fixes appendr :: "'x ⇒ 's ⇒ 's"
assumes appendr_correct:
"invar s ⟹ α (appendr x s) = α s @ [x]"
"invar s ⟹ invar (appendr x s)"
begin
abbreviation (input) "enqueue ≡ appendr"
lemmas enqueue_correct = appendr_correct
end
locale list_remover = list +
constrains α :: "'s ⇒ 'x list"
fixes remover :: "'s ⇒ 's × 'x"
assumes remover_correct:
"⟦invar s; α s ≠ []⟧ ⟹ α (fst (remover s)) = butlast (α s)"
"⟦invar s; α s ≠ []⟧ ⟹ snd (remover s) = last (α s)"
"⟦invar s; α s ≠ []⟧ ⟹ invar (fst (remover s))"
locale list_rightmost = list +
constrains α :: "'s ⇒ 'x list"
fixes rightmost :: "'s ⇒ 'x"
assumes rightmost_correct:
"⟦invar s; α s ≠ []⟧ ⟹ rightmost s = List.last (α s)"
begin
abbreviation (input) bot where "bot ≡ rightmost"
lemmas bot_correct = rightmost_correct
end
subsubsection "Indexing"
locale list_get = list +
constrains α :: "'s ⇒ 'x list"
fixes get :: "'s ⇒ nat ⇒ 'x"
assumes get_correct:
"⟦invar s; i<length (α s)⟧ ⟹ get s i = α s ! i"
locale list_set = list +
constrains α :: "'s ⇒ 'x list"
fixes set :: "'s ⇒ nat ⇒ 'x ⇒ 's"
assumes set_correct:
"⟦invar s; i<length (α s)⟧ ⟹ α (set s i x) = (α s) [i := x]"
"⟦invar s; i<length (α s)⟧ ⟹ invar (set s i x)"