Theory Predicate_Compile_Quickcheck_Examples

Up to index of Isabelle/HOL/Predicate_Compile_Examples

theory Predicate_Compile_Quickcheck_Examples
imports Predicate_Compile_Quickcheck

theory Predicate_Compile_Quickcheck_Examples
imports Predicate_Compile_Quickcheck
begin


section {* Sets *}

lemma "x ∈ {(1::nat)} ==> False"
quickcheck[generator=predicate_compile_wo_ff, iterations=10]
oops

lemma "x ∈ {Suc 0, Suc (Suc 0)} ==> x ≠ Suc 0"
quickcheck[generator=predicate_compile_wo_ff]
oops

lemma "x ∈ {Suc 0, Suc (Suc 0)} ==> x = Suc 0"
quickcheck[generator=predicate_compile_wo_ff]
oops

lemma "x ∈ {Suc 0, Suc (Suc 0)} ==> x <= Suc 0"
quickcheck[generator=predicate_compile_wo_ff]
oops

section {* Numerals *}

lemma
"x ∈ {1, 2, (3::nat)} ==> x = 1 ∨ x = 2"

quickcheck[generator=predicate_compile_wo_ff]
oops

lemma "x ∈ {1, 2, (3::nat)} ==> x < 3"
quickcheck[generator=predicate_compile_wo_ff]
oops

lemma
"x ∈ {1, 2} ∪ {3, 4} ==> x = (1::nat) ∨ x = (2::nat)"

quickcheck[generator=predicate_compile_wo_ff]
oops

section {* Context Free Grammar *}

datatype alphabet = a | b

inductive_set S1 and A1 and B1 where
"[] ∈ S1"
| "w ∈ A1 ==> b # w ∈ S1"
| "w ∈ B1 ==> a # w ∈ S1"
| "w ∈ S1 ==> a # w ∈ A1"
| "w ∈ S1 ==> b # w ∈ S1"
| "[|v ∈ B1; v ∈ B1|] ==> a # v @ w ∈ B1"


lemma
"w ∈ S1 ==> w = []"

quickcheck[generator = predicate_compile_ff_nofs, iterations=1]
oops

theorem S1_sound:
"w ∈ S1 ==> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"

quickcheck[generator=predicate_compile_ff_nofs, size=15]
oops


inductive_set S2 and A2 and B2 where
"[] ∈ S2"
| "w ∈ A2 ==> b # w ∈ S2"
| "w ∈ B2 ==> a # w ∈ S2"
| "w ∈ S2 ==> a # w ∈ A2"
| "w ∈ S2 ==> b # w ∈ B2"
| "[|v ∈ B2; v ∈ B2|] ==> a # v @ w ∈ B2"

(*
code_pred [random_dseq inductify] S2 .
thm S2.random_dseq_equation
thm A2.random_dseq_equation
thm B2.random_dseq_equation

values [random_dseq 1, 2, 8] 10 "{x. S2 x}"

lemma "w ∈ S2 ==> w ≠ [] ==> w ≠ [b, a] ==> w ∈ {}"
quickcheck[generator=predicate_compile, size=8]
oops

lemma "[x <- w. x = a] = []"
quickcheck[generator=predicate_compile]
oops

declare list.size(3,4)[code_pred_def]

(*
lemma "length ([x \<leftarrow> w. x = a]) = (0::nat)"
quickcheck[generator=predicate_compile]
oops
*)

lemma
"w ∈ S2 ==> length [x \<leftarrow> w. x = a] <= Suc (Suc 0)"
quickcheck[generator=predicate_compile, size = 10, iterations = 1]
oops
*)

theorem S2_sound:
"w ∈ S2 --> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"

quickcheck[generator=predicate_compile_ff_nofs, size=5, iterations=10]
oops

inductive_set S3 and A3 and B3 where
"[] ∈ S3"
| "w ∈ A3 ==> b # w ∈ S3"
| "w ∈ B3 ==> a # w ∈ S3"
| "w ∈ S3 ==> a # w ∈ A3"
| "w ∈ S3 ==> b # w ∈ B3"
| "[|v ∈ B3; w ∈ B3|] ==> a # v @ w ∈ B3"


code_pred [inductify, skip_proof] S3 .
thm S3.equation
(*
values 10 "{x. S3 x}"
*)



lemma S3_sound:
"w ∈ S3 --> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"

quickcheck[generator=predicate_compile_ff_fs, size=10, iterations=10]
oops

lemma "¬ (length w > 2) ∨ ¬ (length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b])"
quickcheck[size=10, generator = predicate_compile_ff_fs]
oops

theorem S3_complete:
"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. b = x] --> w ∈ S3"

(*quickcheck[generator=SML]*)
quickcheck[generator=predicate_compile_ff_fs, size=10, iterations=100]
oops


inductive_set S4 and A4 and B4 where
"[] ∈ S4"
| "w ∈ A4 ==> b # w ∈ S4"
| "w ∈ B4 ==> a # w ∈ S4"
| "w ∈ S4 ==> a # w ∈ A4"
| "[|v ∈ A4; w ∈ A4|] ==> b # v @ w ∈ A4"
| "w ∈ S4 ==> b # w ∈ B4"
| "[|v ∈ B4; w ∈ B4|] ==> a # v @ w ∈ B4"


theorem S4_sound:
"w ∈ S4 --> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"

quickcheck[generator = predicate_compile_ff_nofs, size=5, iterations=1]
oops

theorem S4_complete:
"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] --> w ∈ S4"

quickcheck[generator = predicate_compile_ff_nofs, size=5, iterations=1]
oops

hide_const a b

subsection {* Lexicographic order *}
(* TODO *)
(*
lemma
"(u, v) : lexord r ==> (x @ u, y @ v) : lexord r"
oops
*)

subsection {* IMP *}

types
var = nat
state = "int list"


datatype com =
Skip |
Ass var "int" |
Seq com com |
IF "state list" com com |
While "state list" com


inductive exec :: "com => state => state => bool" where
"exec Skip s s" |
"exec (Ass x e) s (s[x := e])" |
"exec c1 s1 s2 ==> exec c2 s2 s3 ==> exec (Seq c1 c2) s1 s3" |
"s ∈ set b ==> exec c1 s t ==> exec (IF b c1 c2) s t" |
"s ∉ set b ==> exec c2 s t ==> exec (IF b c1 c2) s t" |
"s ∉ set b ==> exec (While b c) s s" |
"s1 ∈ set b ==> exec c s1 s2 ==> exec (While b c) s2 s3 ==> exec (While b c) s1 s3"


code_pred [random_dseq] exec .

values [random_dseq 1, 2, 3] 10 "{(c, s, s'). exec c s s'}"

lemma
"exec c s s' ==> exec (Seq c c) s s'"

(*quickcheck[generator = predicate_compile_wo_ff, size=2, iterations=10]*)
oops

subsection {* Lambda *}

datatype type =
Atom nat
| Fun type type (infixr "=>" 200)


datatype dB =
Var nat
| App dB dB (infixl "°" 200)
| Abs type dB


primrec
nth_el :: "'a list => nat => 'a option" ("_⟨_⟩" [90, 0] 91)
where
"[]⟨i⟩ = None"
| "(x # xs)⟨i⟩ = (case i of 0 => Some x | Suc j => xs ⟨j⟩)"


inductive nth_el' :: "'a list => nat => 'a => bool"
where
"nth_el' (x # xs) 0 x"
| "nth_el' xs i y ==> nth_el' (x # xs) (Suc i) y"


inductive typing :: "type list => dB => type => bool" ("_ \<turnstile> _ : _" [50, 50, 50] 50)
where
Var [intro!]: "nth_el' env x T ==> env \<turnstile> Var x : T"
| Abs [intro!]: "T # env \<turnstile> t : U ==> env \<turnstile> Abs T t : (T => U)"
| App [intro!]: "env \<turnstile> s : U => T ==> env \<turnstile> t : T ==> env \<turnstile> (s ° t) : U"


primrec
lift :: "[dB, nat] => dB"
where
"lift (Var i) k = (if i < k then Var i else Var (i + 1))"
| "lift (s ° t) k = lift s k ° lift t k"
| "lift (Abs T s) k = Abs T (lift s (k + 1))"


primrec
subst :: "[dB, dB, nat] => dB" ("_[_'/_]" [300, 0, 0] 300)
where
subst_Var: "(Var i)[s/k] =
(if k < i then Var (i - 1) else if i = k then s else Var i)"

| subst_App: "(t ° u)[s/k] = t[s/k] ° u[s/k]"
| subst_Abs: "(Abs T t)[s/k] = Abs T (t[lift s 0 / k+1])"


inductive beta :: "[dB, dB] => bool" (infixl "->β" 50)
where
beta [simp, intro!]: "Abs T s ° t ->β s[t/0]"
| appL [simp, intro!]: "s ->β t ==> s ° u ->β t ° u"
| appR [simp, intro!]: "s ->β t ==> u ° s ->β u ° t"
| abs [simp, intro!]: "s ->β t ==> Abs T s ->β Abs T t"


lemma
"Γ \<turnstile> t : U ==> t ->β t' ==> Γ \<turnstile> t' : U"

quickcheck[generator = predicate_compile_ff_fs, size = 7, iterations = 10]
oops

subsection {* JAD *}

definition matrix :: "('a :: semiring_0) list list => nat => nat => bool" where
"matrix M rs cs <-> (∀ row ∈ set M. length row = cs) ∧ length M = rs"

(*
code_pred [random_dseq inductify] matrix .
thm matrix.random_dseq_equation

thm matrix_aux.random_dseq_equation

values [random_dseq 3, 2] 10 "{(M, rs, cs). matrix (M:: int list list) rs cs}"
*)

lemma [code_pred_intro]:
"matrix [] 0 m"
"matrix xss n m ==> length xs = m ==> matrix (xs # xss) (Suc n) m"

proof -
show "matrix [] 0 m" unfolding matrix_def by auto
next
show "matrix xss n m ==> length xs = m ==> matrix (xs # xss) (Suc n) m"
unfolding matrix_def by auto
qed

code_pred [random_dseq inductify] matrix
apply (cases x)
unfolding matrix_def apply fastsimp
apply fastsimp done


values [random_dseq 2, 2, 15] 6 "{(M::int list list, n, m). matrix M n m}"

definition "scalar_product v w = (∑ (x, y)\<leftarrow>zip v w. x * y)"

definition mv :: "('a :: semiring_0) list list => 'a list => 'a list"
where [simp]: "mv M v = map (scalar_product v) M"

text {*
This defines the matrix vector multiplication. To work properly @{term
"matrix M m n ∧ length v = n"} must hold.
*}


subsection "Compressed matrix"

definition "sparsify xs = [i \<leftarrow> zip [0..<length xs] xs. snd i ≠ 0]"
(*
lemma sparsify_length: "(i, x) ∈ set (sparsify xs) ==> i < length xs"
by (auto simp: sparsify_def set_zip)

lemma listsum_sparsify[simp]:
fixes v :: "('a :: semiring_0) list"
assumes "length w = length v"
shows "(∑x\<leftarrow>sparsify w. (λ(i, x). v ! i) x * snd x) = scalar_product v w"
(is "(∑x\<leftarrow>_. ?f x) = _")
unfolding sparsify_def scalar_product_def
using assms listsum_map_filter[where f="?f" and P="λ i. snd i ≠ (0::'a)"]
by (simp add: listsum_setsum)
*)

definition [simp]: "unzip w = (map fst w, map snd w)"

primrec insert :: "('a => 'b :: linorder) => 'a => 'a list => 'a list" where
"insert f x [] = [x]" |
"insert f x (y # ys) = (if f y < f x then y # insert f x ys else x # y # ys)"


primrec sort :: "('a => 'b :: linorder) => 'a list => 'a list" where
"sort f [] = []" |
"sort f (x # xs) = insert f x (sort f xs)"


definition
"length_permutate M = (unzip o sort (length o snd)) (zip [0 ..< length M] M)"

(*
definition
"transpose M = [map (λ xs. xs ! i) (takeWhile (λ xs. i < length xs) M). i \<leftarrow> [0 ..< length (M ! 0)]]"
*)

definition
"inflate upds = foldr (λ (i, x) upds. upds[i := x]) upds (replicate (length upds) 0)"


definition
"jad = apsnd transpose o length_permutate o map sparsify"


definition
"jad_mv v = inflate o split zip o apsnd (map listsum o transpose o map (map (λ (i, x). v ! i * x)))"


lemma "matrix (M::int list list) rs cs ==> False"
quickcheck[generator = predicate_compile_ff_nofs, size = 6]
oops

lemma
"[| matrix M rs cs ; length v = cs |] ==> jad_mv v (jad M) = mv M v"

quickcheck[generator = predicate_compile_wo_ff]
oops

end