Theory More_List

Up to index of Isabelle/HOL/MicroJava

theory More_List
imports Main

(*  Author:  Florian Haftmann, TU Muenchen *)

header {* Operations on lists beyond the standard List theory *}

theory More_List
imports Main
begin


hide_const (open) Finite_Set.fold

text {* Repairing code generator setup *}

declare (in lattice) Inf_fin_set_fold [code_unfold del]
declare (in lattice) Sup_fin_set_fold [code_unfold del]
declare (in linorder) Min_fin_set_fold [code_unfold del]
declare (in linorder) Max_fin_set_fold [code_unfold del]
declare (in complete_lattice) Inf_set_fold [code_unfold del]
declare (in complete_lattice) Sup_set_fold [code_unfold del]
declare rev_foldl_cons [code del]

text {* Fold combinator with canonical argument order *}

primrec fold :: "('a => 'b => 'b) => 'a list => 'b => 'b" where
"fold f [] = id"
| "fold f (x # xs) = fold f xs o f x"


lemma foldl_fold:
"foldl f s xs = fold (λx s. f s x) xs s"

by (induct xs arbitrary: s) simp_all

lemma foldr_fold_rev:
"foldr f xs = fold f (rev xs)"

by (simp add: foldr_foldl foldl_fold expand_fun_eq)

lemma fold_rev_conv [code_unfold]:
"fold f (rev xs) = foldr f xs"

by (simp add: foldr_fold_rev)

lemma fold_cong [fundef_cong, recdef_cong]:
"a = b ==> xs = ys ==> (!!x. x ∈ set xs ==> f x = g x)
==> fold f xs a = fold g ys b"

by (induct ys arbitrary: a b xs) simp_all

lemma fold_id:
assumes "!!x. x ∈ set xs ==> f x = id"
shows "fold f xs = id"

using assms by (induct xs) simp_all

lemma fold_apply:
assumes "!!x. x ∈ set xs ==> h o g x = f x o h"
shows "h o fold g xs = fold f xs o h"

using assms by (induct xs) (simp_all add: expand_fun_eq)

lemma fold_invariant:
assumes "!!x. x ∈ set xs ==> Q x" and "P s"
and "!!x s. Q x ==> P s ==> P (f x s)"
shows "P (fold f xs s)"

using assms by (induct xs arbitrary: s) simp_all

lemma fold_weak_invariant:
assumes "P s"
and "!!s x. x ∈ set xs ==> P s ==> P (f x s)"
shows "P (fold f xs s)"

using assms by (induct xs arbitrary: s) simp_all

lemma fold_append [simp]:
"fold f (xs @ ys) = fold f ys o fold f xs"

by (induct xs) simp_all

lemma fold_map [code_unfold]:
"fold g (map f xs) = fold (g o f) xs"

by (induct xs) simp_all

lemma fold_rev:
assumes "!!x y. x ∈ set xs ==> y ∈ set xs ==> f y o f x = f x o f y"
shows "fold f (rev xs) = fold f xs"

using assms by (induct xs) (simp_all del: o_apply add: fold_apply)

lemma foldr_fold:
assumes "!!x y. x ∈ set xs ==> y ∈ set xs ==> f y o f x = f x o f y"
shows "foldr f xs = fold f xs"

using assms unfolding foldr_fold_rev by (rule fold_rev)

lemma fold_Cons_rev:
"fold Cons xs = append (rev xs)"

by (induct xs) simp_all

lemma rev_conv_fold [code]:
"rev xs = fold Cons xs []"

by (simp add: fold_Cons_rev)

lemma fold_append_concat_rev:
"fold append xss = append (concat (rev xss))"

by (induct xss) simp_all

lemma concat_conv_foldr [code]:
"concat xss = foldr append xss []"

by (simp add: fold_append_concat_rev foldr_fold_rev)

lemma fold_plus_listsum_rev:
"fold plus xs = plus (listsum (rev xs))"

by (induct xs) (simp_all add: add.assoc)

lemma listsum_conv_foldr [code]:
"listsum xs = foldr plus xs 0"

by (fact listsum_foldr)

lemma sort_key_conv_fold:
assumes "inj_on f (set xs)"
shows "sort_key f xs = fold (insort_key f) xs []"

proof -
have "fold (insort_key f) (rev xs) = fold (insort_key f) xs"
proof (rule fold_rev, rule ext)
fix zs
fix x y
assume "x ∈ set xs" "y ∈ set xs"
with assms have *: "f y = f x ==> y = x" by (auto dest: inj_onD)
show "(insort_key f y o insort_key f x) zs = (insort_key f x o insort_key f y) zs"
by (induct zs) (auto dest: *)
qed
then show ?thesis by (simp add: sort_key_def foldr_fold_rev)
qed

lemma sort_conv_fold:
"sort xs = fold insort xs []"

by (rule sort_key_conv_fold) simp

text {* @{const Finite_Set.fold} and @{const fold} *}

lemma (in fun_left_comm) fold_set_remdups:
"Finite_Set.fold f y (set xs) = fold f (remdups xs) y"

by (rule sym, induct xs arbitrary: y) (simp_all add: fold_fun_comm insert_absorb)

lemma (in fun_left_comm_idem) fold_set:
"Finite_Set.fold f y (set xs) = fold f xs y"

by (rule sym, induct xs arbitrary: y) (simp_all add: fold_fun_comm)

lemma (in ab_semigroup_idem_mult) fold1_set:
assumes "xs ≠ []"
shows "Finite_Set.fold1 times (set xs) = fold times (tl xs) (hd xs)"

proof -
interpret fun_left_comm_idem times by (fact fun_left_comm_idem)
from assms obtain y ys where xs: "xs = y # ys"
by (cases xs) auto
show ?thesis
proof (cases "set ys = {}")
case True with xs show ?thesis by simp
next
case False
then have "fold1 times (insert y (set ys)) = Finite_Set.fold times y (set ys)"
by (simp only: finite_set fold1_eq_fold_idem)
with xs show ?thesis by (simp add: fold_set mult_commute)
qed
qed

lemma (in lattice) Inf_fin_set_fold:
"Inf_fin (set (x # xs)) = fold inf xs x"

proof -
interpret ab_semigroup_idem_mult "inf :: 'a => 'a => 'a"
by (fact ab_semigroup_idem_mult_inf)
show ?thesis
by (simp add: Inf_fin_def fold1_set del: set.simps)
qed

lemma (in lattice) Inf_fin_set_foldr [code_unfold]:
"Inf_fin (set (x # xs)) = foldr inf xs x"

by (simp add: Inf_fin_set_fold ac_simps foldr_fold expand_fun_eq del: set.simps)

lemma (in lattice) Sup_fin_set_fold:
"Sup_fin (set (x # xs)) = fold sup xs x"

proof -
interpret ab_semigroup_idem_mult "sup :: 'a => 'a => 'a"
by (fact ab_semigroup_idem_mult_sup)
show ?thesis
by (simp add: Sup_fin_def fold1_set del: set.simps)
qed

lemma (in lattice) Sup_fin_set_foldr [code_unfold]:
"Sup_fin (set (x # xs)) = foldr sup xs x"

by (simp add: Sup_fin_set_fold ac_simps foldr_fold expand_fun_eq del: set.simps)

lemma (in linorder) Min_fin_set_fold:
"Min (set (x # xs)) = fold min xs x"

proof -
interpret ab_semigroup_idem_mult "min :: 'a => 'a => 'a"
by (fact ab_semigroup_idem_mult_min)
show ?thesis
by (simp add: Min_def fold1_set del: set.simps)
qed

lemma (in linorder) Min_fin_set_foldr [code_unfold]:
"Min (set (x # xs)) = foldr min xs x"

by (simp add: Min_fin_set_fold ac_simps foldr_fold expand_fun_eq del: set.simps)

lemma (in linorder) Max_fin_set_fold:
"Max (set (x # xs)) = fold max xs x"

proof -
interpret ab_semigroup_idem_mult "max :: 'a => 'a => 'a"
by (fact ab_semigroup_idem_mult_max)
show ?thesis
by (simp add: Max_def fold1_set del: set.simps)
qed

lemma (in linorder) Max_fin_set_foldr [code_unfold]:
"Max (set (x # xs)) = foldr max xs x"

by (simp add: Max_fin_set_fold ac_simps foldr_fold expand_fun_eq del: set.simps)

lemma (in complete_lattice) Inf_set_fold:
"Inf (set xs) = fold inf xs top"

proof -
interpret fun_left_comm_idem "inf :: 'a => 'a => 'a"
by (fact fun_left_comm_idem_inf)
show ?thesis by (simp add: Inf_fold_inf fold_set inf_commute)
qed

lemma (in complete_lattice) Inf_set_foldr [code_unfold]:
"Inf (set xs) = foldr inf xs top"

by (simp add: Inf_set_fold ac_simps foldr_fold expand_fun_eq)

lemma (in complete_lattice) Sup_set_fold:
"Sup (set xs) = fold sup xs bot"

proof -
interpret fun_left_comm_idem "sup :: 'a => 'a => 'a"
by (fact fun_left_comm_idem_sup)
show ?thesis by (simp add: Sup_fold_sup fold_set sup_commute)
qed

lemma (in complete_lattice) Sup_set_foldr [code_unfold]:
"Sup (set xs) = foldr sup xs bot"

by (simp add: Sup_set_fold ac_simps foldr_fold expand_fun_eq)

lemma (in complete_lattice) INFI_set_fold:
"INFI (set xs) f = fold (inf o f) xs top"

unfolding INFI_def set_map [symmetric] Inf_set_fold fold_map ..

lemma (in complete_lattice) SUPR_set_fold:
"SUPR (set xs) f = fold (sup o f) xs bot"

unfolding SUPR_def set_map [symmetric] Sup_set_fold fold_map ..

text {* @{text nth_map} *}

definition nth_map :: "nat => ('a => 'a) => 'a list => 'a list" where
"nth_map n f xs = (if n < length xs then
take n xs @ [f (xs ! n)] @ drop (Suc n) xs
else xs)"


lemma nth_map_id:
"n ≥ length xs ==> nth_map n f xs = xs"

by (simp add: nth_map_def)

lemma nth_map_unfold:
"n < length xs ==> nth_map n f xs = take n xs @ [f (xs ! n)] @ drop (Suc n) xs"

by (simp add: nth_map_def)

lemma nth_map_Nil [simp]:
"nth_map n f [] = []"

by (simp add: nth_map_def)

lemma nth_map_zero [simp]:
"nth_map 0 f (x # xs) = f x # xs"

by (simp add: nth_map_def)

lemma nth_map_Suc [simp]:
"nth_map (Suc n) f (x # xs) = x # nth_map n f xs"

by (simp add: nth_map_def)

end