lib: add defn of list_insert_before, and some lemmas

Signed-off-by: Michael McInerney <michael.mcinerney@proofcraft.systems>
This commit is contained in:
Michael McInerney 2024-01-10 10:54:38 +10:30 committed by Achim D. Brucker
parent f857b158a1
commit 74d8f98e85
2 changed files with 32 additions and 0 deletions

View File

@ -371,6 +371,32 @@ lemma list_insert_after_after:
apply fastforce
done
lemma list_insert_before_not_found:
"a \<notin> set ls \<Longrightarrow> list_insert_before ls a new = ls"
by (induct ls; fastforce)
lemma list_insert_before_nonempty:
"ls \<noteq> [] \<Longrightarrow> list_insert_before ls a new \<noteq> []"
by (induct ls; clarsimp)
lemma list_insert_before_head:
"xs \<noteq> [] \<Longrightarrow> list_insert_before xs (hd xs) new = new # xs"
by (induct xs; fastforce)
lemma last_of_list_insert_before:
"xs \<noteq> [] \<Longrightarrow> last (list_insert_before xs a new) = last xs"
by (induct xs; clarsimp simp: list_insert_before_nonempty)
lemma list_insert_before_distinct:
"\<lbrakk>distinct (xs @ ys); ys \<noteq> []\<rbrakk> \<Longrightarrow> list_insert_before (xs @ ys) (hd ys) new = xs @ new # ys"
by (induct xs; fastforce simp add: list_insert_before_head)
lemma set_list_insert_before:
"\<lbrakk>new \<notin> set ls; before \<in> set ls\<rbrakk> \<Longrightarrow> set (list_insert_before ls before new) = set ls \<union> {new}"
apply (induct ls; clarsimp)
apply fastforce
done
lemma list_remove_removed:
"set (list_remove list x) = (set list) - {x}"
apply (induct list,simp+)

View File

@ -26,6 +26,12 @@ primrec list_insert_after :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a \<Righta
"list_insert_after (x # xs) a b = (if x = a then x # b # xs
else x # list_insert_after xs a b)"
\<comment> \<open>Inserts the value new immediately before the first occurence of a (if any) in the list\<close>
primrec list_insert_before :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a list" where
"list_insert_before [] a new = []" |
"list_insert_before (x # xs) a new = (if x = a then new # x # xs
else x # list_insert_before xs a new)"
primrec list_remove :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a list" where
"list_remove [] a = []" |
"list_remove (x # xs) a = (if x = a then (list_remove xs a)