Theory Specialisation_Examples
theory Specialisation_Examples
imports Main "HOL-Library.Predicate_Compile_Alternative_Defs"
begin
declare [[values_timeout = 960.0]]
section ‹Specialisation Examples›
primrec nth_el'
where
"nth_el' [] i = None"
| "nth_el' (x # xs) i = (case i of 0 => Some x | Suc j => nth_el' xs j)"
definition
"greater_than_index xs = (∀i x. nth_el' xs i = Some x --> x > i)"
code_pred (expected_modes: i => bool) [inductify, skip_proof, specialise] greater_than_index .
ML_val ‹Core_Data.intros_of \<^context> \<^const_name>‹specialised_nth_el'P››
thm greater_than_index.equation
values [expected "{()}"] "{x. greater_than_index [1,2,4,6]}"
values [expected "{}"] "{x. greater_than_index [0,2,3,2]}"
subsection ‹Common subterms›
text ‹If a predicate is called with common subterms as arguments,
this predicate should be specialised.
›
definition max_nat :: "nat => nat => nat"
where "max_nat a b = (if a <= b then b else a)"
lemma [code_pred_inline]:
"max = max_nat"
by (simp add: fun_eq_iff max_def max_nat_def)
definition
"max_of_my_Suc x = max x (Suc x)"
text ‹In this example, max is specialised, hence the mode o => i => bool is possible›
code_pred (modes: o => i => bool) [inductify, specialise, skip_proof] max_of_my_Suc .
thm max_of_my_SucP.equation
ML_val ‹Core_Data.intros_of \<^context> \<^const_name>‹specialised_max_natP››
values "{x. max_of_my_SucP x 6}"
subsection ‹Sorts›
inductive sorted :: "'a::linorder list ⇒ bool" where
Nil [simp]: "sorted []"
| Cons: "∀y∈set xs. x ≤ y ⟹ sorted xs ⟹ sorted (x # xs)"
lemma sorted_single [simp]: "sorted [x]"
by (rule sorted.Cons) auto
lemma sorted_many: "x ≤ y ⟹ sorted (y # zs) ⟹ sorted (x # y # zs)"
by (rule sorted.Cons) (cases "y # zs" rule: sorted.cases, auto)
lemma sorted_many_eq [simp]:
"sorted (x # y # zs) ⟷ x ≤ y ∧ sorted (y # zs)"
by (auto intro: sorted_many elim: sorted.cases)
declare sorted.Nil [code_pred_intro]
sorted_single [code_pred_intro]
sorted_many [code_pred_intro]
code_pred sorted
proof -
assume "sorted xa"
assume 1: "xa = [] ⟹ thesis"
assume 2: "⋀x. xa = [x] ⟹ thesis"
assume 3: "⋀x y zs. xa = x # y # zs ⟹ x ≤ y ⟹ sorted (y # zs) ⟹ thesis"
show thesis proof (cases xa)
case Nil with 1 show ?thesis .
next
case (Cons x xs) show ?thesis proof (cases xs)
case Nil with Cons 2 show ?thesis by simp
next
case (Cons y zs)
show ?thesis
proof (rule 3)
from Cons ‹xa = x # xs› show "xa = x # y # zs" by simp
with ‹sorted xa› show "x ≤ y" and "sorted (y # zs)" by simp_all
qed
qed
qed
qed
thm sorted.equation
section ‹Specialisation in POPLmark theory›
notation
Some ("⌊_⌋")
notation
None ("⊥")
notation
length ("∥_∥")
notation
Cons ("_ ∷/ _" [66, 65] 65)
primrec
nth_el :: "'a list ⇒ nat ⇒ 'a option" ("_⟨_⟩" [90, 0] 91)
where
"[]⟨i⟩ = ⊥"
| "(x # xs)⟨i⟩ = (case i of 0 ⇒ ⌊x⌋ | Suc j ⇒ xs ⟨j⟩)"
primrec assoc :: "('a × 'b) list ⇒ 'a ⇒ 'b option" ("_⟨_⟩⇩?" [90, 0] 91)
where
"[]⟨a⟩⇩? = ⊥"
| "(x # xs)⟨a⟩⇩? = (if fst x = a then ⌊snd x⌋ else xs⟨a⟩⇩?)"
primrec unique :: "('a × 'b) list ⇒ bool"
where
"unique [] = True"
| "unique (x # xs) = (xs⟨fst x⟩⇩? = ⊥ ∧ unique xs)"